aboutsummaryrefslogtreecommitdiff
path: root/prover_snapshots/coq/RV64
diff options
context:
space:
mode:
Diffstat (limited to 'prover_snapshots/coq/RV64')
-rw-r--r--prover_snapshots/coq/RV64/mem_metadata.v8
-rw-r--r--prover_snapshots/coq/RV64/riscv.v82833
-rw-r--r--prover_snapshots/coq/RV64/riscv_extras.v32
-rw-r--r--prover_snapshots/coq/RV64/riscv_types.v14096
4 files changed, 55131 insertions, 41838 deletions
diff --git a/prover_snapshots/coq/RV64/mem_metadata.v b/prover_snapshots/coq/RV64/mem_metadata.v
new file mode 100644
index 0000000..bbfc856
--- /dev/null
+++ b/prover_snapshots/coq/RV64/mem_metadata.v
@@ -0,0 +1,8 @@
+Require Import Sail.Base.
+Open Scope Z.
+
+Definition write_ram {rv e a} wk (addr : mword a) size (v : mword (8 * size)) (meta : unit) : monad rv bool e := write_mem wk a addr size v.
+
+Definition read_ram {rv e a} rk (addr : mword a) size (read_tag : bool) `{ArithFact (size >=? 0)} : monad rv (mword (8 * size) * unit) e :=
+ read_mem rk a addr size >>= fun data =>
+ returnm (data, tt).
diff --git a/prover_snapshots/coq/RV64/riscv.v b/prover_snapshots/coq/RV64/riscv.v
index fc0e61d..57c4eaa 100644
--- a/prover_snapshots/coq/RV64/riscv.v
+++ b/prover_snapshots/coq/RV64/riscv.v
@@ -1,387 +1,347 @@
(*Generated by Sail from riscv.*)
-Require Import Sail2_instr_kinds.
-Require Import Sail2_values.
-Require Import Sail2_string.
-Require Import Sail2_real.
-Require Import Sail2_operators_mwords.
-Require Import Sail2_prompt_monad.
-Require Import Sail2_prompt.
-Require Import Sail2_state.
+Require Import Sail.Base.
+Require Import Sail.Real.
Require Import riscv_types.
+Require Import mem_metadata.
Require Import riscv_extras.
Import ListNotations.
Open Scope string.
Open Scope bool.
-Section Content.
+Open Scope Z.
-Definition is_none {a : Type} (opt : option a)
-: bool :=
-
- match opt with | Some (_) => false | None => true end.
-Definition is_some {a : Type} (opt : option a)
-: bool :=
-
- match opt with | Some (_) => true | None => false end.
+Definition is_none {a : Type} (opt : option a) : bool :=
+ match opt with | Some _ => false | None => true end.
+
+Definition is_some {a : Type} (opt : option a) : bool :=
+ match opt with | Some _ => true | None => false end.
+
+Definition eq_unit (_ : unit) (_ : unit) : {_bool : bool & ArithFact (_bool)} := build_ex (true).
+
+Definition neq_int (x : Z) (y : Z) : {_bool : bool & ArithFact (Bool.eqb (negb (x =? y)) _bool)} :=
+ build_ex (negb (Z.eqb x y)).
-Definition eq_unit (_ : unit) (_ : unit)
-: {_bool : bool & ArithFact (_bool = true)} :=
-
- build_ex(true).
+Definition neq_bool (x : bool) (y : bool) : bool := negb (Bool.eqb x y).
-Definition neq_int (x : Z) (y : Z)
-: {_bool : bool & ArithFact (iff (_bool = true) (x <> y))} :=
-
- build_ex(negb (Z.eqb x y)).
+Definition __id (x : Z) : {_retval : Z & ArithFact (_retval =? x)} := build_ex (x).
-Definition neq_bool (x : bool) (y : bool) : bool := negb (Bool.eqb x y).
+Definition fdiv_int (n : Z) (m : Z) : Z :=
+ if sumbool_of_bool (andb (Z.ltb n 0) (Z.gtb m 0)) then Z.sub (Z.quot (Z.add n 1) m) 1
+ else if sumbool_of_bool (andb (Z.gtb n 0) (Z.ltb m 0)) then Z.sub (Z.quot (Z.sub n 1) m) 1
+ else Z.quot n m.
-Definition __id (x : Z) : {_retval : Z & ArithFact (_retval = x)} := build_ex(x).
+Definition fmod_int (n : Z) (m : Z) : Z := Z.sub n (Z.mul m (fdiv_int n m)).
-Definition concat_str_bits {n : Z} (str : string) (x : mword n)
-: string :=
-
+Definition concat_str_bits {n : Z} (str : string) (x : mword n) : string :=
String.append str (string_of_bits x).
-Definition concat_str_dec (str : string) (x : Z) : string := String.append str (dec_str x).
+Definition concat_str_dec (str : string) (x : Z) : string := String.append str (dec_str x).
-Definition sail_mask {v0 : Z} (len : Z) (v : mword v0) `{ArithFact (len >= 0 /\ v0 >= 0)}
+Definition sail_mask {v0 : Z} (len : Z) (v : mword v0) `{ArithFact ((len >=? 0) && (v0 >=? 0))}
: mword len :=
-
- if sumbool_of_bool ((Z.leb len (length_mword v))) then vector_truncate v len
- else zero_extend v len.
+ if sumbool_of_bool (Z.leb len (length_mword v)) then vector_truncate v len else zero_extend v len.
-Definition sail_ones (n : Z) `{ArithFact (n >= 0)} : mword n := not_vec (zeros n).
+Definition sail_ones (n : Z) `{ArithFact (n >=? 0)} : mword n := not_vec (zeros n).
-Definition slice_mask (n : Z) (i : Z) (l : Z) `{ArithFact (n >= 0)}
-: mword n :=
-
- if sumbool_of_bool ((Z.geb l n)) then shiftl (sail_ones n) i
+Definition slice_mask (n : Z) (i : Z) (l : Z) `{ArithFact (n >=? 0)} : mword n :=
+ if sumbool_of_bool (Z.geb l n) then shiftl (sail_ones n) i
else
- let one : bits n := sail_mask n ((vec_of_bits [B1] : mword 1) : bits 1) in
+ let one : bits n := sail_mask n ('b"1" : bits 1) in
shiftl (sub_vec (shiftl one l) one) i.
-Definition read_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 11)}
-: read_kind :=
-
- let l__196 := arg_ in
- if sumbool_of_bool ((Z.eqb l__196 0)) then Read_plain
- else if sumbool_of_bool ((Z.eqb l__196 1)) then Read_reserve
- else if sumbool_of_bool ((Z.eqb l__196 2)) then Read_acquire
- else if sumbool_of_bool ((Z.eqb l__196 3)) then Read_exclusive
- else if sumbool_of_bool ((Z.eqb l__196 4)) then Read_exclusive_acquire
- else if sumbool_of_bool ((Z.eqb l__196 5)) then Read_stream
- else if sumbool_of_bool ((Z.eqb l__196 6)) then Read_RISCV_acquire
- else if sumbool_of_bool ((Z.eqb l__196 7)) then Read_RISCV_strong_acquire
- else if sumbool_of_bool ((Z.eqb l__196 8)) then Read_RISCV_reserved
- else if sumbool_of_bool ((Z.eqb l__196 9)) then Read_RISCV_reserved_acquire
- else if sumbool_of_bool ((Z.eqb l__196 10)) then Read_RISCV_reserved_strong_acquire
+Definition read_kind_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 11))} : read_kind :=
+ let l__292 := arg_ in
+ if sumbool_of_bool (Z.eqb l__292 0) then Read_plain
+ else if sumbool_of_bool (Z.eqb l__292 1) then Read_reserve
+ else if sumbool_of_bool (Z.eqb l__292 2) then Read_acquire
+ else if sumbool_of_bool (Z.eqb l__292 3) then Read_exclusive
+ else if sumbool_of_bool (Z.eqb l__292 4) then Read_exclusive_acquire
+ else if sumbool_of_bool (Z.eqb l__292 5) then Read_stream
+ else if sumbool_of_bool (Z.eqb l__292 6) then Read_RISCV_acquire
+ else if sumbool_of_bool (Z.eqb l__292 7) then Read_RISCV_strong_acquire
+ else if sumbool_of_bool (Z.eqb l__292 8) then Read_RISCV_reserved
+ else if sumbool_of_bool (Z.eqb l__292 9) then Read_RISCV_reserved_acquire
+ else if sumbool_of_bool (Z.eqb l__292 10) then Read_RISCV_reserved_strong_acquire
else Read_X86_locked.
-Definition num_of_read_kind (arg_ : read_kind)
-: {e : Z & ArithFact (0 <= e /\ e <= 11)} :=
-
- build_ex(match arg_ with
- | Read_plain => 0
- | Read_reserve => 1
- | Read_acquire => 2
- | Read_exclusive => 3
- | Read_exclusive_acquire => 4
- | Read_stream => 5
- | Read_RISCV_acquire => 6
- | Read_RISCV_strong_acquire => 7
- | Read_RISCV_reserved => 8
- | Read_RISCV_reserved_acquire => 9
- | Read_RISCV_reserved_strong_acquire => 10
- | Read_X86_locked => 11
- end).
-
-Definition write_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 10)}
-: write_kind :=
-
- let l__186 := arg_ in
- if sumbool_of_bool ((Z.eqb l__186 0)) then Write_plain
- else if sumbool_of_bool ((Z.eqb l__186 1)) then Write_conditional
- else if sumbool_of_bool ((Z.eqb l__186 2)) then Write_release
- else if sumbool_of_bool ((Z.eqb l__186 3)) then Write_exclusive
- else if sumbool_of_bool ((Z.eqb l__186 4)) then Write_exclusive_release
- else if sumbool_of_bool ((Z.eqb l__186 5)) then Write_RISCV_release
- else if sumbool_of_bool ((Z.eqb l__186 6)) then Write_RISCV_strong_release
- else if sumbool_of_bool ((Z.eqb l__186 7)) then Write_RISCV_conditional
- else if sumbool_of_bool ((Z.eqb l__186 8)) then Write_RISCV_conditional_release
- else if sumbool_of_bool ((Z.eqb l__186 9)) then Write_RISCV_conditional_strong_release
+Definition num_of_read_kind (arg_ : read_kind) : {e : Z & ArithFact ((0 <=? e) && (e <=? 11))} :=
+ build_ex (
+ match arg_ with
+ | Read_plain => 0
+ | Read_reserve => 1
+ | Read_acquire => 2
+ | Read_exclusive => 3
+ | Read_exclusive_acquire => 4
+ | Read_stream => 5
+ | Read_RISCV_acquire => 6
+ | Read_RISCV_strong_acquire => 7
+ | Read_RISCV_reserved => 8
+ | Read_RISCV_reserved_acquire => 9
+ | Read_RISCV_reserved_strong_acquire => 10
+ | Read_X86_locked => 11
+ end
+ ).
+
+Definition write_kind_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 10))} : write_kind :=
+ let l__282 := arg_ in
+ if sumbool_of_bool (Z.eqb l__282 0) then Write_plain
+ else if sumbool_of_bool (Z.eqb l__282 1) then Write_conditional
+ else if sumbool_of_bool (Z.eqb l__282 2) then Write_release
+ else if sumbool_of_bool (Z.eqb l__282 3) then Write_exclusive
+ else if sumbool_of_bool (Z.eqb l__282 4) then Write_exclusive_release
+ else if sumbool_of_bool (Z.eqb l__282 5) then Write_RISCV_release
+ else if sumbool_of_bool (Z.eqb l__282 6) then Write_RISCV_strong_release
+ else if sumbool_of_bool (Z.eqb l__282 7) then Write_RISCV_conditional
+ else if sumbool_of_bool (Z.eqb l__282 8) then Write_RISCV_conditional_release
+ else if sumbool_of_bool (Z.eqb l__282 9) then Write_RISCV_conditional_strong_release
else Write_X86_locked.
-Definition num_of_write_kind (arg_ : write_kind)
-: {e : Z & ArithFact (0 <= e /\ e <= 10)} :=
-
- build_ex(match arg_ with
- | Write_plain => 0
- | Write_conditional => 1
- | Write_release => 2
- | Write_exclusive => 3
- | Write_exclusive_release => 4
- | Write_RISCV_release => 5
- | Write_RISCV_strong_release => 6
- | Write_RISCV_conditional => 7
- | Write_RISCV_conditional_release => 8
- | Write_RISCV_conditional_strong_release => 9
- | Write_X86_locked => 10
- end).
-
-Definition a64_barrier_domain_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
+Definition num_of_write_kind (arg_ : write_kind) : {e : Z & ArithFact ((0 <=? e) && (e <=? 10))} :=
+ build_ex (
+ match arg_ with
+ | Write_plain => 0
+ | Write_conditional => 1
+ | Write_release => 2
+ | Write_exclusive => 3
+ | Write_exclusive_release => 4
+ | Write_RISCV_release => 5
+ | Write_RISCV_strong_release => 6
+ | Write_RISCV_conditional => 7
+ | Write_RISCV_conditional_release => 8
+ | Write_RISCV_conditional_strong_release => 9
+ | Write_X86_locked => 10
+ end
+ ).
+
+Definition a64_barrier_domain_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))}
: a64_barrier_domain :=
-
- let l__183 := arg_ in
- if sumbool_of_bool ((Z.eqb l__183 0)) then A64_FullShare
- else if sumbool_of_bool ((Z.eqb l__183 1)) then A64_InnerShare
- else if sumbool_of_bool ((Z.eqb l__183 2)) then A64_OuterShare
+ let l__279 := arg_ in
+ if sumbool_of_bool (Z.eqb l__279 0) then A64_FullShare
+ else if sumbool_of_bool (Z.eqb l__279 1) then A64_InnerShare
+ else if sumbool_of_bool (Z.eqb l__279 2) then A64_OuterShare
else A64_NonShare.
-Definition num_of_a64_barrier_domain (arg_ : a64_barrier_domain)
-: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
-
- build_ex(match arg_ with
- | A64_FullShare => 0
- | A64_InnerShare => 1
- | A64_OuterShare => 2
- | A64_NonShare => 3
- end).
-
-Definition a64_barrier_type_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
+Definition num_of_a64_barrier_domain (arg_ : a64_barrier_domain)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (
+ match arg_ with
+ | A64_FullShare => 0
+ | A64_InnerShare => 1
+ | A64_OuterShare => 2
+ | A64_NonShare => 3
+ end
+ ).
+
+Definition a64_barrier_type_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))}
: a64_barrier_type :=
-
- let l__181 := arg_ in
- if sumbool_of_bool ((Z.eqb l__181 0)) then A64_barrier_all
- else if sumbool_of_bool ((Z.eqb l__181 1)) then A64_barrier_LD
+ let l__277 := arg_ in
+ if sumbool_of_bool (Z.eqb l__277 0) then A64_barrier_all
+ else if sumbool_of_bool (Z.eqb l__277 1) then A64_barrier_LD
else A64_barrier_ST.
-Definition num_of_a64_barrier_type (arg_ : a64_barrier_type)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | A64_barrier_all => 0 | A64_barrier_LD => 1 | A64_barrier_ST => 2 end).
-
-Definition trans_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
-: trans_kind :=
-
- let l__179 := arg_ in
- if sumbool_of_bool ((Z.eqb l__179 0)) then Transaction_start
- else if sumbool_of_bool ((Z.eqb l__179 1)) then Transaction_commit
+Definition num_of_a64_barrier_type (arg_ : a64_barrier_type)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | A64_barrier_all => 0 | A64_barrier_LD => 1 | A64_barrier_ST => 2 end).
+
+Definition trans_kind_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : trans_kind :=
+ let l__275 := arg_ in
+ if sumbool_of_bool (Z.eqb l__275 0) then Transaction_start
+ else if sumbool_of_bool (Z.eqb l__275 1) then Transaction_commit
else Transaction_abort.
-Definition num_of_trans_kind (arg_ : trans_kind)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with
- | Transaction_start => 0
- | Transaction_commit => 1
- | Transaction_abort => 2
- end).
+Definition num_of_trans_kind (arg_ : trans_kind) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (
+ match arg_ with
+ | Transaction_start => 0
+ | Transaction_commit => 1
+ | Transaction_abort => 2
+ end
+ ).
-Definition cache_op_kind_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 10)}
+Definition cache_op_kind_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 10))}
: cache_op_kind :=
-
- let l__169 := arg_ in
- if sumbool_of_bool ((Z.eqb l__169 0)) then Cache_op_D_IVAC
- else if sumbool_of_bool ((Z.eqb l__169 1)) then Cache_op_D_ISW
- else if sumbool_of_bool ((Z.eqb l__169 2)) then Cache_op_D_CSW
- else if sumbool_of_bool ((Z.eqb l__169 3)) then Cache_op_D_CISW
- else if sumbool_of_bool ((Z.eqb l__169 4)) then Cache_op_D_ZVA
- else if sumbool_of_bool ((Z.eqb l__169 5)) then Cache_op_D_CVAC
- else if sumbool_of_bool ((Z.eqb l__169 6)) then Cache_op_D_CVAU
- else if sumbool_of_bool ((Z.eqb l__169 7)) then Cache_op_D_CIVAC
- else if sumbool_of_bool ((Z.eqb l__169 8)) then Cache_op_I_IALLUIS
- else if sumbool_of_bool ((Z.eqb l__169 9)) then Cache_op_I_IALLU
+ let l__265 := arg_ in
+ if sumbool_of_bool (Z.eqb l__265 0) then Cache_op_D_IVAC
+ else if sumbool_of_bool (Z.eqb l__265 1) then Cache_op_D_ISW
+ else if sumbool_of_bool (Z.eqb l__265 2) then Cache_op_D_CSW
+ else if sumbool_of_bool (Z.eqb l__265 3) then Cache_op_D_CISW
+ else if sumbool_of_bool (Z.eqb l__265 4) then Cache_op_D_ZVA
+ else if sumbool_of_bool (Z.eqb l__265 5) then Cache_op_D_CVAC
+ else if sumbool_of_bool (Z.eqb l__265 6) then Cache_op_D_CVAU
+ else if sumbool_of_bool (Z.eqb l__265 7) then Cache_op_D_CIVAC
+ else if sumbool_of_bool (Z.eqb l__265 8) then Cache_op_I_IALLUIS
+ else if sumbool_of_bool (Z.eqb l__265 9) then Cache_op_I_IALLU
else Cache_op_I_IVAU.
-Definition num_of_cache_op_kind (arg_ : cache_op_kind)
-: {e : Z & ArithFact (0 <= e /\ e <= 10)} :=
-
- build_ex(match arg_ with
- | Cache_op_D_IVAC => 0
- | Cache_op_D_ISW => 1
- | Cache_op_D_CSW => 2
- | Cache_op_D_CISW => 3
- | Cache_op_D_ZVA => 4
- | Cache_op_D_CVAC => 5
- | Cache_op_D_CVAU => 6
- | Cache_op_D_CIVAC => 7
- | Cache_op_I_IALLUIS => 8
- | Cache_op_I_IALLU => 9
- | Cache_op_I_IVAU => 10
- end).
+Definition num_of_cache_op_kind (arg_ : cache_op_kind)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 10))} :=
+ build_ex (
+ match arg_ with
+ | Cache_op_D_IVAC => 0
+ | Cache_op_D_ISW => 1
+ | Cache_op_D_CSW => 2
+ | Cache_op_D_CISW => 3
+ | Cache_op_D_ZVA => 4
+ | Cache_op_D_CVAC => 5
+ | Cache_op_D_CVAU => 6
+ | Cache_op_D_CIVAC => 7
+ | Cache_op_I_IALLUIS => 8
+ | Cache_op_I_IALLU => 9
+ | Cache_op_I_IVAU => 10
+ end
+ ).
-Definition neq_vec {n : Z} (x : mword n) (y : mword n) : bool := negb (eq_vec x y).
+Definition not_bit (b : bitU) : bitU := if eq_bit b B1 then B0 else B1.
+Definition neq_vec {n : Z} (x : mword n) (y : mword n) : bool := negb (eq_vec x y).
-Definition cast_unit_vec (b : bitU)
-: M (mword 1) :=
-
- (match b with
- | B0 => returnm ((vec_of_bits [B0] : mword 1) : mword 1)
- | B1 => returnm ((vec_of_bits [B1] : mword 1) : mword 1)
- | _ => exit tt : M (mword 1)
- end)
- : M (mword 1).
-Definition get_config_print_instr '(tt : unit) : bool := false.
+Definition string_of_bit (b : bitU) : M (string) :=
+ (match b with | B0 => returnm "0b0" | B1 => returnm "0b1" | _ => exit tt : M (string) end)
+ : M (string).
+
+Definition get_config_print_instr '(tt : unit) : bool := false.
-Definition get_config_print_reg '(tt : unit) : bool := false.
+Definition get_config_print_reg '(tt : unit) : bool := false.
-Definition get_config_print_mem '(tt : unit) : bool := false.
+Definition get_config_print_mem '(tt : unit) : bool := false.
-Definition get_config_print_platform '(tt : unit) : bool := false.
+Definition get_config_print_platform '(tt : unit) : bool := false.
-Definition EXTS {n : Z} (m : Z) (v : mword n) `{ArithFact (m >= n)} : mword m := sign_extend v m.
+Definition EXTS {n : Z} (m : Z) (v : mword n) `{ArithFact (m >=? n)} : mword m := sign_extend v m.
-Definition EXTZ {n : Z} (m : Z) (v : mword n) `{ArithFact (m >= n)} : mword m := zero_extend v m.
+Definition EXTZ {n : Z} (m : Z) (v : mword n) `{ArithFact (m >=? n)} : mword m := zero_extend v m.
-Definition zeros_implicit (n : Z) `{ArithFact (n >= 0)} : mword n := zeros n.
+Definition zeros_implicit (n : Z) `{ArithFact (n >=? 0)} : mword n := zeros n.
-Definition ones (n : Z) `{ArithFact (n >= 0)} : mword n := sail_ones n.
+Definition ones (n : Z) `{ArithFact (n >=? 0)} : mword n := sail_ones n.
-Definition bool_to_bits (x : bool)
-: mword 1 :=
-
- if sumbool_of_bool (x) then (vec_of_bits [B1] : mword 1)
- else (vec_of_bits [B0] : mword 1).
+Definition bool_to_bits (x : bool) : mword 1 :=
+ if sumbool_of_bool x then 'b"1" : mword 1 else 'b"0" : mword 1.
-Definition bit_to_bool (b : bitU)
-: M (bool) :=
-
- (match b with
- | B1 => returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | B0 => returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))}))
- | _ => exit tt : M (bool)
- end)
+Definition bit_to_bool (b : bitU) : M (bool) :=
+ (match b with | B1 => returnm true | B0 => returnm false | _ => exit tt : M (bool) end)
: M (bool).
-Definition to_bits (l : Z) (n : Z) `{ArithFact (l >= 0)} : mword l := get_slice_int l n 0.
+Definition to_bits (l : Z) (n : Z) `{ArithFact (l >=? 0)} : mword l := get_slice_int l n 0.
-Definition zopz0zI_s {n : Z} (x : mword n) (y : mword n) `{ArithFact (n > 0)}
-: bool :=
-
+Definition zopz0zI_s {n : Z} (x : mword n) (y : mword n) `{ArithFact (n >? 0)} : bool :=
Z.ltb (projT1 (sint x)) (projT1 (sint y)).
-Definition zopz0zKzJ_s {n : Z} (x : mword n) (y : mword n) `{ArithFact (n > 0)}
-: bool :=
-
+Definition zopz0zKzJ_s {n : Z} (x : mword n) (y : mword n) `{ArithFact (n >? 0)} : bool :=
Z.geb (projT1 (sint x)) (projT1 (sint y)).
-Definition zopz0zI_u {n : Z} (x : mword n) (y : mword n)
-: bool :=
-
+Definition zopz0zI_u {n : Z} (x : mword n) (y : mword n) : bool :=
Z.ltb (projT1 (uint x)) (projT1 (uint y)).
-Definition zopz0zKzJ_u {n : Z} (x : mword n) (y : mword n)
-: bool :=
-
+Definition zopz0zKzJ_u {n : Z} (x : mword n) (y : mword n) : bool :=
Z.geb (projT1 (uint x)) (projT1 (uint y)).
-Definition zopz0zIzJ_u {n : Z} (x : mword n) (y : mword n)
-: bool :=
-
+Definition zopz0zIzJ_u {n : Z} (x : mword n) (y : mword n) : bool :=
Z.leb (projT1 (uint x)) (projT1 (uint y)).
-Definition shift_right_arith64 (v : mword 64) (shift : mword 6)
-: mword 64 :=
-
+Definition shift_right_arith64 (v : mword 64) (shift : mword 6) : mword 64 :=
let v128 : bits 128 := EXTS 128 v in
subrange_vec_dec (shift_bits_right v128 shift) 63 0.
-Definition shift_right_arith32 (v : mword 32) (shift : mword 5)
-: mword 32 :=
-
+Definition shift_right_arith32 (v : mword 32) (shift : mword 5) : mword 32 :=
let v64 : bits 64 := EXTS 64 v in
subrange_vec_dec (shift_bits_right v64 shift) 31 0.
-Fixpoint _rec_n_leading_spaces (s : string) (_reclimit : Z) (_acc : Acc (Zwf 0) _reclimit)
-{struct _acc} : M ({n : Z & ArithFact (n >= 0)}) :=
-
+Axiom spc_forwards_matches : forall (_ : unit) , bool.
+
+Axiom spc_backwards_matches : forall (_ : string) , bool.
+
+Axiom opt_spc_forwards_matches : forall (_ : unit) , bool.
+
+Axiom opt_spc_backwards_matches : forall (_ : string) , bool.
+
+Axiom def_spc_forwards_matches : forall (_ : unit) , bool.
+
+Axiom def_spc_backwards_matches : forall (_ : string) , bool.
+
+Axiom hex_bits_forwards : forall {n : Z} (_ : (Z * mword n)) , string.
+
+Axiom hex_bits_backwards : forall {n : Z} (_ : string) , (Z * mword n).
+
+Axiom hex_bits_forwards_matches : forall {n : Z} (_ : (Z * mword n)) , bool.
+
+Axiom hex_bits_backwards_matches : forall (_ : string) , bool.
+
+Axiom hex_bits_matches_prefix : forall
+{n : Z}
+(_ : string)
+,
+option (((Z * mword n) * {n : Z & ArithFact (n >=? 0)})).
+
+Fixpoint _rec_n_leading_spaces (s : string) (_reclimit : Z) (_acc : Acc (Zwf 0) _reclimit)
+{struct _acc} : M ({n : Z & ArithFact (n >=? 0)}).
+exact (
assert_exp' (Z.geb _reclimit 0) "recursion limit reached" >>= fun _ =>
let p0_ := s in
- (if ((generic_eq p0_ "")) then returnm (build_ex (0 : Z))
+ (if generic_eq p0_ "" then returnm (build_ex 0)
else
let p0_ := string_take s 1 in
- (if ((generic_eq p0_ " ")) then
- (_rec_n_leading_spaces (string_drop s 1) (Z.sub _reclimit 1) (_limit_reduces _acc)) >>= fun '(existT _ w__0 _ : {n : Z & ArithFact (n >=
+ (if generic_eq p0_ " " then
+ (_rec_n_leading_spaces (string_drop s 1) (Z.sub _reclimit 1) (_limit_reduces _acc)) >>= fun '(existT _ w__0 _ : {n : Z & ArithFact (n >=?
0)}) =>
- returnm (build_ex
- (projT1
- (build_ex
- (Z.add 1 w__0)
- : {_atom : Z & ArithFact (exists ex98922_ , _atom = (1 + ex98922_) /\ 0 <= ex98922_)})))
- else returnm (build_ex (0 : Z)))
- : M ({n : Z & ArithFact (n >= 0)}))
- : M ({n : Z & ArithFact (n >= 0)}).
-
-Definition n_leading_spaces (s : string)
-: M ({n : Z & ArithFact (n >= 0)}) :=
-
+ returnm (build_ex (Z.add 1 w__0))
+ else returnm (build_ex 0))
+ : M ({n : Z & ArithFact (n >=? 0)}))
+ : M ({n : Z & ArithFact (n >=? 0)})
+).
+Defined.
+
+
+Definition n_leading_spaces (s : string) : M ({n : Z & ArithFact (n >=? 0)}) :=
(_rec_n_leading_spaces s ((projT1 (string_length s)) : Z) (Zwf_guarded _))
- : M ({n : Z & ArithFact (n >= 0)}).
+ : M ({n : Z & ArithFact (n >=? 0)}).
-Definition spc_forwards '(tt : unit) : string := " ".
+Definition spc_forwards '(tt : unit) : string := " ".
-Definition spc_backwards (s : string) : unit := tt.
+Definition spc_backwards (s : string) : unit := tt.
-Definition spc_matches_prefix (s : string)
-: M (option ((unit * {n : Z & ArithFact (n >= 0)}))) :=
-
+Definition spc_matches_prefix (s : string) : M (option ((unit * {n : Z & ArithFact (n >=? 0)}))) :=
(n_leading_spaces s) >>= fun '(existT _ n _) =>
- let l__168 := n in
- returnm ((if sumbool_of_bool ((Z.eqb l__168 0)) then None
- else Some ((tt, build_ex n)))
- : option ((unit * {n : Z & ArithFact (n >= 0)}))).
+ let l__264 := n in
+ returnm (if sumbool_of_bool (Z.eqb l__264 0) then None else Some (tt, build_ex n)).
-Definition opt_spc_forwards '(tt : unit) : string := "".
+Definition opt_spc_forwards '(tt : unit) : string := "".
-Definition opt_spc_backwards (s : string) : unit := tt.
+Definition opt_spc_backwards (s : string) : unit := tt.
-Definition opt_spc_matches_prefix (s : string)
-: M (option ((unit * {n : Z & ArithFact (n >= 0)}))) :=
-
- (n_leading_spaces s) >>= fun '(existT _ w__0 _ : {n : Z & ArithFact (n >= 0)}) =>
- returnm ((Some
- ((tt, build_ex
- w__0)))
- : option ((unit * {n : Z & ArithFact (n >= 0)}))).
+Definition opt_spc_matches_prefix (s : string) : M (option ((unit * {n : Z & ArithFact (n >=? 0)}))) :=
+ (n_leading_spaces s) >>= fun '(existT _ w__0 _ : {n : Z & ArithFact (n >=? 0)}) =>
+ returnm (Some (tt, build_ex w__0)).
-Definition def_spc_forwards '(tt : unit) : string := " ".
+Definition def_spc_forwards '(tt : unit) : string := " ".
-Definition def_spc_backwards (s : string) : unit := tt.
+Definition def_spc_backwards (s : string) : unit := tt.
-Definition def_spc_matches_prefix (s : string)
-: M (option ((unit * {n : Z & ArithFact (n >= 0)}))) :=
-
- (opt_spc_matches_prefix s)
- : M (option ((unit * {n : Z & ArithFact (n >= 0)}))).
+Definition def_spc_matches_prefix (s : string) : M (option ((unit * {n : Z & ArithFact (n >=? 0)}))) :=
+ (opt_spc_matches_prefix s) : M (option ((unit * {n : Z & ArithFact (n >=? 0)}))).
-Definition hex_bits_1_forwards_matches (bv : mword 1) : bool := true.
+Definition hex_bits_1_forwards_matches (bv : mword 1) : bool := true.
-Definition hex_bits_1_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_1_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_1_matches_prefix s) with
- | Some ((g__258, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_1_matches_prefix s) with
+ | Some (g__361, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_1_backwards (s : string)
-: M (mword 1) :=
-
+Definition hex_bits_1_backwards (s : string) : M (mword 1) :=
(match (hex_bits_1_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 1)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 49:2 - 51:3" >>= fun _ =>
exit tt)
@@ -392,28 +352,24 @@ Definition hex_bits_1_backwards (s : string)
end)
: M (mword 1).
-Definition hex_bits_2_forwards_matches (bv : mword 2) : bool := true.
+Definition hex_bits_2_forwards_matches (bv : mword 2) : bool := true.
-Definition hex_bits_2_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_2_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_2_matches_prefix s) with
- | Some ((g__257, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_2_matches_prefix s) with
+ | Some (g__360, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_2_backwards (s : string)
-: M (mword 2) :=
-
+Definition hex_bits_2_backwards (s : string) : M (mword 2) :=
(match (hex_bits_2_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 2)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 68:2 - 70:3" >>= fun _ =>
exit tt)
@@ -424,28 +380,24 @@ Definition hex_bits_2_backwards (s : string)
end)
: M (mword 2).
-Definition hex_bits_3_forwards_matches (bv : mword 3) : bool := true.
+Definition hex_bits_3_forwards_matches (bv : mword 3) : bool := true.
-Definition hex_bits_3_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_3_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_3_matches_prefix s) with
- | Some ((g__256, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_3_matches_prefix s) with
+ | Some (g__359, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_3_backwards (s : string)
-: M (mword 3) :=
-
+Definition hex_bits_3_backwards (s : string) : M (mword 3) :=
(match (hex_bits_3_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 3)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 87:2 - 89:3" >>= fun _ =>
exit tt)
@@ -456,28 +408,24 @@ Definition hex_bits_3_backwards (s : string)
end)
: M (mword 3).
-Definition hex_bits_4_forwards_matches (bv : mword 4) : bool := true.
+Definition hex_bits_4_forwards_matches (bv : mword 4) : bool := true.
-Definition hex_bits_4_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_4_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_4_matches_prefix s) with
- | Some ((g__255, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_4_matches_prefix s) with
+ | Some (g__358, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_4_backwards (s : string)
-: M (mword 4) :=
-
+Definition hex_bits_4_backwards (s : string) : M (mword 4) :=
(match (hex_bits_4_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 4)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 106:2 - 108:3" >>= fun _ =>
exit tt)
@@ -488,28 +436,24 @@ Definition hex_bits_4_backwards (s : string)
end)
: M (mword 4).
-Definition hex_bits_5_forwards_matches (bv : mword 5) : bool := true.
+Definition hex_bits_5_forwards_matches (bv : mword 5) : bool := true.
-Definition hex_bits_5_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_5_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_5_matches_prefix s) with
- | Some ((g__254, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_5_matches_prefix s) with
+ | Some (g__357, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_5_backwards (s : string)
-: M (mword 5) :=
-
+Definition hex_bits_5_backwards (s : string) : M (mword 5) :=
(match (hex_bits_5_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 5)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 125:2 - 127:3" >>= fun _ =>
exit tt)
@@ -520,28 +464,24 @@ Definition hex_bits_5_backwards (s : string)
end)
: M (mword 5).
-Definition hex_bits_6_forwards_matches (bv : mword 6) : bool := true.
+Definition hex_bits_6_forwards_matches (bv : mword 6) : bool := true.
-Definition hex_bits_6_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_6_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_6_matches_prefix s) with
- | Some ((g__253, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_6_matches_prefix s) with
+ | Some (g__356, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_6_backwards (s : string)
-: M (mword 6) :=
-
+Definition hex_bits_6_backwards (s : string) : M (mword 6) :=
(match (hex_bits_6_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 6)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 144:2 - 146:3" >>= fun _ =>
exit tt)
@@ -552,28 +492,24 @@ Definition hex_bits_6_backwards (s : string)
end)
: M (mword 6).
-Definition hex_bits_7_forwards_matches (bv : mword 7) : bool := true.
+Definition hex_bits_7_forwards_matches (bv : mword 7) : bool := true.
-Definition hex_bits_7_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_7_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_7_matches_prefix s) with
- | Some ((g__252, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_7_matches_prefix s) with
+ | Some (g__355, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_7_backwards (s : string)
-: M (mword 7) :=
-
+Definition hex_bits_7_backwards (s : string) : M (mword 7) :=
(match (hex_bits_7_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 7)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 163:2 - 165:3" >>= fun _ =>
exit tt)
@@ -584,28 +520,24 @@ Definition hex_bits_7_backwards (s : string)
end)
: M (mword 7).
-Definition hex_bits_8_forwards_matches (bv : mword 8) : bool := true.
+Definition hex_bits_8_forwards_matches (bv : mword 8) : bool := true.
-Definition hex_bits_8_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_8_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_8_matches_prefix s) with
- | Some ((g__251, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_8_matches_prefix s) with
+ | Some (g__354, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_8_backwards (s : string)
-: M (mword 8) :=
-
+Definition hex_bits_8_backwards (s : string) : M (mword 8) :=
(match (hex_bits_8_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 8)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 182:2 - 184:3" >>= fun _ =>
exit tt)
@@ -616,28 +548,24 @@ Definition hex_bits_8_backwards (s : string)
end)
: M (mword 8).
-Definition hex_bits_9_forwards_matches (bv : mword 9) : bool := true.
+Definition hex_bits_9_forwards_matches (bv : mword 9) : bool := true.
-Definition hex_bits_9_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_9_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_9_matches_prefix s) with
- | Some ((g__250, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_9_matches_prefix s) with
+ | Some (g__353, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_9_backwards (s : string)
-: M (mword 9) :=
-
+Definition hex_bits_9_backwards (s : string) : M (mword 9) :=
(match (hex_bits_9_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 9)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 201:2 - 203:3" >>= fun _ =>
exit tt)
@@ -648,28 +576,24 @@ Definition hex_bits_9_backwards (s : string)
end)
: M (mword 9).
-Definition hex_bits_10_forwards_matches (bv : mword 10) : bool := true.
+Definition hex_bits_10_forwards_matches (bv : mword 10) : bool := true.
-Definition hex_bits_10_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_10_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_10_matches_prefix s) with
- | Some ((g__249, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_10_matches_prefix s) with
+ | Some (g__352, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_10_backwards (s : string)
-: M (mword 10) :=
-
+Definition hex_bits_10_backwards (s : string) : M (mword 10) :=
(match (hex_bits_10_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 10)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 220:2 - 222:3" >>= fun _ =>
exit tt)
@@ -680,28 +604,24 @@ Definition hex_bits_10_backwards (s : string)
end)
: M (mword 10).
-Definition hex_bits_11_forwards_matches (bv : mword 11) : bool := true.
+Definition hex_bits_11_forwards_matches (bv : mword 11) : bool := true.
-Definition hex_bits_11_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_11_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_11_matches_prefix s) with
- | Some ((g__248, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_11_matches_prefix s) with
+ | Some (g__351, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_11_backwards (s : string)
-: M (mword 11) :=
-
+Definition hex_bits_11_backwards (s : string) : M (mword 11) :=
(match (hex_bits_11_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 11)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
assert_exp' false "Pattern match failure at model/prelude_mapping.sail 239:2 - 241:3" >>= fun _ =>
exit tt)
@@ -712,813 +632,712 @@ Definition hex_bits_11_backwards (s : string)
end)
: M (mword 11).
-Definition hex_bits_12_forwards_matches (bv : mword 12) : bool := true.
+Definition hex_bits_12_forwards_matches (bv : mword 12) : bool := true.
-Definition hex_bits_12_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_12_matches_prefix (s : string)
+: option ((mword 12 * {n : Z & ArithFact (n >=? 0)})) :=
+ None.
+
+Definition hex_bits_12_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_12_matches_prefix s) with
- | Some ((g__247, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_12_matches_prefix s) with
+ | Some (g__350, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_12_backwards (s : string)
-: M (mword 12) :=
-
+Definition hex_bits_12_backwards (s : string) : M (mword 12) :=
(match (hex_bits_12_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 12)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 258:2 - 260:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 262:2 - 264:3" >>= fun _ =>
exit tt)
: M (mword 12)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 258:2 - 260:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 262:2 - 264:3" >>= fun _ =>
exit tt
end)
: M (mword 12).
-Definition hex_bits_13_forwards_matches (bv : mword 13) : bool := true.
+Definition hex_bits_13_forwards_matches (bv : mword 13) : bool := true.
-Definition hex_bits_13_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_13_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_13_matches_prefix s) with
- | Some ((g__246, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_13_matches_prefix s) with
+ | Some (g__349, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_13_backwards (s : string)
-: M (mword 13) :=
-
+Definition hex_bits_13_backwards (s : string) : M (mword 13) :=
(match (hex_bits_13_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 13)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 277:2 - 279:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 281:2 - 283:3" >>= fun _ =>
exit tt)
: M (mword 13)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 277:2 - 279:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 281:2 - 283:3" >>= fun _ =>
exit tt
end)
: M (mword 13).
-Definition hex_bits_14_forwards_matches (bv : mword 14) : bool := true.
+Definition hex_bits_14_forwards_matches (bv : mword 14) : bool := true.
-Definition hex_bits_14_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_14_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_14_matches_prefix s) with
- | Some ((g__245, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_14_matches_prefix s) with
+ | Some (g__348, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_14_backwards (s : string)
-: M (mword 14) :=
-
+Definition hex_bits_14_backwards (s : string) : M (mword 14) :=
(match (hex_bits_14_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 14)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 296:2 - 298:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 300:2 - 302:3" >>= fun _ =>
exit tt)
: M (mword 14)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 296:2 - 298:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 300:2 - 302:3" >>= fun _ =>
exit tt
end)
: M (mword 14).
-Definition hex_bits_15_forwards_matches (bv : mword 15) : bool := true.
+Definition hex_bits_15_forwards_matches (bv : mword 15) : bool := true.
-Definition hex_bits_15_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_15_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_15_matches_prefix s) with
- | Some ((g__244, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_15_matches_prefix s) with
+ | Some (g__347, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_15_backwards (s : string)
-: M (mword 15) :=
-
+Definition hex_bits_15_backwards (s : string) : M (mword 15) :=
(match (hex_bits_15_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 15)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 315:2 - 317:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 319:2 - 321:3" >>= fun _ =>
exit tt)
: M (mword 15)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 315:2 - 317:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 319:2 - 321:3" >>= fun _ =>
exit tt
end)
: M (mword 15).
-Definition hex_bits_16_forwards_matches (bv : mword 16) : bool := true.
+Definition hex_bits_16_forwards_matches (bv : mword 16) : bool := true.
-Definition hex_bits_16_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_16_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_16_matches_prefix s) with
- | Some ((g__243, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_16_matches_prefix s) with
+ | Some (g__346, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_16_backwards (s : string)
-: M (mword 16) :=
-
+Definition hex_bits_16_backwards (s : string) : M (mword 16) :=
(match (hex_bits_16_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 16)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 334:2 - 336:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 338:2 - 340:3" >>= fun _ =>
exit tt)
: M (mword 16)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 334:2 - 336:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 338:2 - 340:3" >>= fun _ =>
exit tt
end)
: M (mword 16).
-Definition hex_bits_17_forwards_matches (bv : mword 17) : bool := true.
+Definition hex_bits_17_forwards_matches (bv : mword 17) : bool := true.
-Definition hex_bits_17_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_17_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_17_matches_prefix s) with
- | Some ((g__242, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_17_matches_prefix s) with
+ | Some (g__345, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_17_backwards (s : string)
-: M (mword 17) :=
-
+Definition hex_bits_17_backwards (s : string) : M (mword 17) :=
(match (hex_bits_17_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 17)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 353:2 - 355:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 357:2 - 359:3" >>= fun _ =>
exit tt)
: M (mword 17)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 353:2 - 355:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 357:2 - 359:3" >>= fun _ =>
exit tt
end)
: M (mword 17).
-Definition hex_bits_18_forwards_matches (bv : mword 18) : bool := true.
+Definition hex_bits_18_forwards_matches (bv : mword 18) : bool := true.
-Definition hex_bits_18_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_18_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_18_matches_prefix s) with
- | Some ((g__241, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_18_matches_prefix s) with
+ | Some (g__344, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_18_backwards (s : string)
-: M (mword 18) :=
-
+Definition hex_bits_18_backwards (s : string) : M (mword 18) :=
(match (hex_bits_18_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 18)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 372:2 - 374:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 376:2 - 378:3" >>= fun _ =>
exit tt)
: M (mword 18)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 372:2 - 374:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 376:2 - 378:3" >>= fun _ =>
exit tt
end)
: M (mword 18).
-Definition hex_bits_19_forwards_matches (bv : mword 19) : bool := true.
+Definition hex_bits_19_forwards_matches (bv : mword 19) : bool := true.
-Definition hex_bits_19_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_19_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_19_matches_prefix s) with
- | Some ((g__240, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_19_matches_prefix s) with
+ | Some (g__343, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_19_backwards (s : string)
-: M (mword 19) :=
-
+Definition hex_bits_19_backwards (s : string) : M (mword 19) :=
(match (hex_bits_19_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 19)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 391:2 - 393:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 395:2 - 397:3" >>= fun _ =>
exit tt)
: M (mword 19)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 391:2 - 393:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 395:2 - 397:3" >>= fun _ =>
exit tt
end)
: M (mword 19).
-Definition hex_bits_20_forwards_matches (bv : mword 20) : bool := true.
+Definition hex_bits_20_forwards_matches (bv : mword 20) : bool := true.
-Definition hex_bits_20_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_20_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_20_matches_prefix s) with
- | Some ((g__239, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_20_matches_prefix s) with
+ | Some (g__342, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_20_backwards (s : string)
-: M (mword 20) :=
-
+Definition hex_bits_20_backwards (s : string) : M (mword 20) :=
(match (hex_bits_20_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 20)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 410:2 - 412:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 414:2 - 416:3" >>= fun _ =>
exit tt)
: M (mword 20)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 410:2 - 412:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 414:2 - 416:3" >>= fun _ =>
exit tt
end)
: M (mword 20).
-Definition hex_bits_21_forwards_matches (bv : mword 21) : bool := true.
+Definition hex_bits_21_forwards_matches (bv : mword 21) : bool := true.
-Definition hex_bits_21_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_21_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_21_matches_prefix s) with
- | Some ((g__238, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_21_matches_prefix s) with
+ | Some (g__341, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_21_backwards (s : string)
-: M (mword 21) :=
-
+Definition hex_bits_21_backwards (s : string) : M (mword 21) :=
(match (hex_bits_21_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 21)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 429:2 - 431:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 433:2 - 435:3" >>= fun _ =>
exit tt)
: M (mword 21)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 429:2 - 431:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 433:2 - 435:3" >>= fun _ =>
exit tt
end)
: M (mword 21).
-Definition hex_bits_22_forwards_matches (bv : mword 22) : bool := true.
+Definition hex_bits_22_forwards_matches (bv : mword 22) : bool := true.
-Definition hex_bits_22_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_22_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_22_matches_prefix s) with
- | Some ((g__237, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_22_matches_prefix s) with
+ | Some (g__340, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_22_backwards (s : string)
-: M (mword 22) :=
-
+Definition hex_bits_22_backwards (s : string) : M (mword 22) :=
(match (hex_bits_22_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 22)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 448:2 - 450:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 452:2 - 454:3" >>= fun _ =>
exit tt)
: M (mword 22)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 448:2 - 450:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 452:2 - 454:3" >>= fun _ =>
exit tt
end)
: M (mword 22).
-Definition hex_bits_23_forwards_matches (bv : mword 23) : bool := true.
+Definition hex_bits_23_forwards_matches (bv : mword 23) : bool := true.
-Definition hex_bits_23_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_23_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_23_matches_prefix s) with
- | Some ((g__236, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_23_matches_prefix s) with
+ | Some (g__339, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_23_backwards (s : string)
-: M (mword 23) :=
-
+Definition hex_bits_23_backwards (s : string) : M (mword 23) :=
(match (hex_bits_23_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 23)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 467:2 - 469:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 471:2 - 473:3" >>= fun _ =>
exit tt)
: M (mword 23)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 467:2 - 469:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 471:2 - 473:3" >>= fun _ =>
exit tt
end)
: M (mword 23).
-Definition hex_bits_24_forwards_matches (bv : mword 24) : bool := true.
+Definition hex_bits_24_forwards_matches (bv : mword 24) : bool := true.
-Definition hex_bits_24_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_24_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_24_matches_prefix s) with
- | Some ((g__235, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_24_matches_prefix s) with
+ | Some (g__338, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_24_backwards (s : string)
-: M (mword 24) :=
-
+Definition hex_bits_24_backwards (s : string) : M (mword 24) :=
(match (hex_bits_24_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 24)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 486:2 - 488:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 490:2 - 492:3" >>= fun _ =>
exit tt)
: M (mword 24)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 486:2 - 488:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 490:2 - 492:3" >>= fun _ =>
exit tt
end)
: M (mword 24).
-Definition hex_bits_25_forwards_matches (bv : mword 25) : bool := true.
+Definition hex_bits_25_forwards_matches (bv : mword 25) : bool := true.
-Definition hex_bits_25_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_25_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_25_matches_prefix s) with
- | Some ((g__234, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_25_matches_prefix s) with
+ | Some (g__337, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_25_backwards (s : string)
-: M (mword 25) :=
-
+Definition hex_bits_25_backwards (s : string) : M (mword 25) :=
(match (hex_bits_25_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 25)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 505:2 - 507:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 509:2 - 511:3" >>= fun _ =>
exit tt)
: M (mword 25)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 505:2 - 507:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 509:2 - 511:3" >>= fun _ =>
exit tt
end)
: M (mword 25).
-Definition hex_bits_26_forwards_matches (bv : mword 26) : bool := true.
+Definition hex_bits_26_forwards_matches (bv : mword 26) : bool := true.
-Definition hex_bits_26_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_26_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_26_matches_prefix s) with
- | Some ((g__233, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_26_matches_prefix s) with
+ | Some (g__336, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_26_backwards (s : string)
-: M (mword 26) :=
-
+Definition hex_bits_26_backwards (s : string) : M (mword 26) :=
(match (hex_bits_26_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 26)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 524:2 - 526:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 528:2 - 530:3" >>= fun _ =>
exit tt)
: M (mword 26)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 524:2 - 526:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 528:2 - 530:3" >>= fun _ =>
exit tt
end)
: M (mword 26).
-Definition hex_bits_27_forwards_matches (bv : mword 27) : bool := true.
+Definition hex_bits_27_forwards_matches (bv : mword 27) : bool := true.
-Definition hex_bits_27_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_27_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_27_matches_prefix s) with
- | Some ((g__232, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_27_matches_prefix s) with
+ | Some (g__335, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_27_backwards (s : string)
-: M (mword 27) :=
-
+Definition hex_bits_27_backwards (s : string) : M (mword 27) :=
(match (hex_bits_27_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 27)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 543:2 - 545:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 547:2 - 549:3" >>= fun _ =>
exit tt)
: M (mword 27)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 543:2 - 545:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 547:2 - 549:3" >>= fun _ =>
exit tt
end)
: M (mword 27).
-Definition hex_bits_28_forwards_matches (bv : mword 28) : bool := true.
+Definition hex_bits_28_forwards_matches (bv : mword 28) : bool := true.
-Definition hex_bits_28_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_28_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_28_matches_prefix s) with
- | Some ((g__231, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_28_matches_prefix s) with
+ | Some (g__334, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_28_backwards (s : string)
-: M (mword 28) :=
-
+Definition hex_bits_28_backwards (s : string) : M (mword 28) :=
(match (hex_bits_28_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 28)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 562:2 - 564:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 566:2 - 568:3" >>= fun _ =>
exit tt)
: M (mword 28)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 562:2 - 564:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 566:2 - 568:3" >>= fun _ =>
exit tt
end)
: M (mword 28).
-Definition hex_bits_29_forwards_matches (bv : mword 29) : bool := true.
+Definition hex_bits_29_forwards_matches (bv : mword 29) : bool := true.
-Definition hex_bits_29_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_29_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_29_matches_prefix s) with
- | Some ((g__230, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_29_matches_prefix s) with
+ | Some (g__333, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_29_backwards (s : string)
-: M (mword 29) :=
-
+Definition hex_bits_29_backwards (s : string) : M (mword 29) :=
(match (hex_bits_29_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 29)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 581:2 - 583:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 585:2 - 587:3" >>= fun _ =>
exit tt)
: M (mword 29)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 581:2 - 583:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 585:2 - 587:3" >>= fun _ =>
exit tt
end)
: M (mword 29).
-Definition hex_bits_30_forwards_matches (bv : mword 30) : bool := true.
+Definition hex_bits_30_forwards_matches (bv : mword 30) : bool := true.
-Definition hex_bits_30_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_30_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_30_matches_prefix s) with
- | Some ((g__229, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_30_matches_prefix s) with
+ | Some (g__332, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_30_backwards (s : string)
-: M (mword 30) :=
-
+Definition hex_bits_30_backwards (s : string) : M (mword 30) :=
(match (hex_bits_30_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 30)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 600:2 - 602:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 604:2 - 606:3" >>= fun _ =>
exit tt)
: M (mword 30)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 600:2 - 602:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 604:2 - 606:3" >>= fun _ =>
exit tt
end)
: M (mword 30).
-Definition hex_bits_31_forwards_matches (bv : mword 31) : bool := true.
+Definition hex_bits_31_forwards_matches (bv : mword 31) : bool := true.
-Definition hex_bits_31_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_31_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_31_matches_prefix s) with
- | Some ((g__228, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_31_matches_prefix s) with
+ | Some (g__331, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_31_backwards (s : string)
-: M (mword 31) :=
-
+Definition hex_bits_31_backwards (s : string) : M (mword 31) :=
(match (hex_bits_31_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 31)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 619:2 - 621:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 623:2 - 625:3" >>= fun _ =>
exit tt)
: M (mword 31)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 619:2 - 621:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 623:2 - 625:3" >>= fun _ =>
exit tt
end)
: M (mword 31).
-Definition hex_bits_32_forwards_matches (bv : mword 32) : bool := true.
+Definition hex_bits_32_forwards_matches (bv : mword 32) : bool := true.
-Definition hex_bits_32_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_32_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_32_matches_prefix s) with
- | Some ((g__227, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_32_matches_prefix s) with
+ | Some (g__330, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_32_backwards (s : string)
-: M (mword 32) :=
-
+Definition hex_bits_32_backwards (s : string) : M (mword 32) :=
(match (hex_bits_32_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 32)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 638:2 - 640:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 642:2 - 644:3" >>= fun _ =>
exit tt)
: M (mword 32)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 638:2 - 640:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 642:2 - 644:3" >>= fun _ =>
exit tt
end)
: M (mword 32).
-Definition hex_bits_33_forwards_matches (bv : mword 33) : bool := true.
+Definition hex_bits_33_forwards_matches (bv : mword 33) : bool := true.
-Definition hex_bits_33_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_33_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_33_matches_prefix s) with
- | Some ((g__226, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_33_matches_prefix s) with
+ | Some (g__329, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_33_backwards (s : string)
-: M (mword 33) :=
-
+Definition hex_bits_33_backwards (s : string) : M (mword 33) :=
(match (hex_bits_33_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 33)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 657:2 - 659:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 661:2 - 663:3" >>= fun _ =>
exit tt)
: M (mword 33)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 657:2 - 659:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 661:2 - 663:3" >>= fun _ =>
exit tt
end)
: M (mword 33).
-Definition hex_bits_48_forwards_matches (bv : mword 48) : bool := true.
+Definition hex_bits_48_forwards_matches (bv : mword 48) : bool := true.
-Definition hex_bits_48_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_48_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_48_matches_prefix s) with
- | Some ((g__225, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_48_matches_prefix s) with
+ | Some (g__328, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_48_backwards (s : string)
-: M (mword 48) :=
-
+Definition hex_bits_48_backwards (s : string) : M (mword 48) :=
(match (hex_bits_48_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 48)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 676:2 - 678:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 680:2 - 682:3" >>= fun _ =>
exit tt)
: M (mword 48)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 676:2 - 678:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 680:2 - 682:3" >>= fun _ =>
exit tt
end)
: M (mword 48).
-Definition hex_bits_64_forwards_matches (bv : mword 64) : bool := true.
+Definition hex_bits_64_forwards_matches (bv : mword 64) : bool := true.
-Definition hex_bits_64_backwards_matches (s : string)
-: bool :=
-
+Definition hex_bits_64_backwards_matches (s : string) : bool :=
match s with
| s =>
- if ((match (hex_bits_64_matches_prefix s) with
- | Some ((g__224, existT _ n _)) =>
- if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then true else false
- | _ => false
- end)) then
+ if match (hex_bits_64_matches_prefix s) with
+ | Some (g__327, existT _ n _) =>
+ if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then true else false
+ | _ => false
+ end then
true
else false
end.
-Definition hex_bits_64_backwards (s : string)
-: M (mword 64) :=
-
+Definition hex_bits_64_backwards (s : string) : M (mword 64) :=
(match (hex_bits_64_matches_prefix s) with
- | Some ((bv, existT _ n _)) =>
- (if sumbool_of_bool ((Z.eqb n (projT1 (string_length s)))) then returnm (bv : mword 64)
+ | Some (bv, existT _ n _) =>
+ (if sumbool_of_bool (Z.eqb n (projT1 (string_length s))) then returnm bv
else
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 695:2 - 697:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 699:2 - 701:3" >>= fun _ =>
exit tt)
: M (mword 64)
| _ =>
- assert_exp' false "Pattern match failure at model/prelude_mapping.sail 695:2 - 697:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/prelude_mapping.sail 699:2 - 701:3" >>= fun _ =>
exit tt
end)
: M (mword 64).
Definition default_meta : mem_meta := tt.
Hint Unfold default_meta : sail.
-Definition __WriteRAM_Meta (addr : mword 64) (width : Z) (meta : unit)
-: M (unit) :=
-
- returnm (tt
- : unit).
-
-Definition __ReadRAM_Meta (addr : mword 64) (width : Z) : M (unit) := returnm (tt : unit).
-
-Definition write_ram
-(wk : write_kind) (addr : mword 64) (width : Z) (data : mword (8 * width)) (meta : unit)
-`{ArithFact (0 < width /\ width <= 16)}
-: M (bool) :=
-
- (write_mem wk 64 addr width data) >>= fun ret : bool =>
- (if sumbool_of_bool (ret) then (__WriteRAM_Meta addr width meta) : M (unit)
- else returnm (tt : unit)) >>
- returnm (ret
- : bool).
+Definition __WriteRAM_Meta (addr : mword 64) (width : Z) (meta : unit) : M (unit) := returnm tt.
+
+Definition __ReadRAM_Meta (addr : mword 64) (width : Z) : M (unit) := returnm tt.
+
+
Definition write_ram_ea (wk : write_kind) (addr : mword 64) (width : Z)
-`{ArithFact (0 < width /\ width <= 16)}
+`{ArithFact ((0 <? width) && (width <=? 16))}
: M (unit) :=
-
- (write_mem_ea wk 64 addr width)
- : M (unit).
+ (write_mem_ea wk 64 addr width) : M (unit).
+
-Definition read_ram (rk : read_kind) (addr : mword 64) (width : Z)
-`{ArithFact (0 < width /\ width <= 16)}
-: M (mword (8 * width)) :=
-
- (read_mem rk 64 addr width)
- : M (mword (8 * width)).
Axiom __TraceMemoryWrite : forall {m : Z} (n : Z) (_ : mword m) (_ : mword (8 * n)) , unit.
Axiom __TraceMemoryRead : forall {m : Z} (n : Z) (_ : mword m) (_ : mword (8 * n)) , unit.
+Definition init_ext_ptw : ext_ptw := tt.
+Hint Unfold init_ext_ptw : sail.
+Definition ext_translate_exception (e : unit) : unit := e.
+
+Definition ext_exc_type_to_bits (e : unit) : mword 8 := Ox"18" : mword 8.
+
+Definition num_of_ext_exc_type (e : unit) : {n : Z & ArithFact ((0 <=? n) && (n <? xlen))} :=
+ build_ex (24).
+
+Definition ext_exc_type_to_str (e : unit) : string := "extension-exception".
+
Definition xlen_val := 64.
Hint Unfold xlen_val : sail.
Definition xlen_max_unsigned := Z.sub (projT1 (pow2 64)) 1.
@@ -1527,2188 +1346,2023 @@ Definition xlen_max_signed := Z.sub (projT1 (pow2 (Z.sub 64 1))) 1.
Hint Unfold xlen_max_signed : sail.
Definition xlen_min_signed := Z.sub 0 (projT1 (pow2 (Z.sub 64 1))).
Hint Unfold xlen_min_signed : sail.
-Definition regidx_to_regno (b : mword 5)
-: {n : Z & ArithFact (0 <= n /\ n < 32)} :=
-
- build_ex(let 'r := projT1 (uint b) in
- r).
-
-Definition creg2reg_idx (creg : mword 3)
-: mword 5 :=
-
- concat_vec (vec_of_bits [B0;B1] : mword 2) creg.
-
-Definition zreg : regidx := (vec_of_bits [B0;B0;B0;B0;B0] : mword 5).
+Definition regidx_to_regno (b : mword 5) : {n : Z & ArithFact ((0 <=? n) && (n <? 32))} :=
+ build_ex (
+ let r := projT1 (uint b) in
+ r
+ ).
+
+Definition creg2reg_idx (creg : mword 3) : mword 5 := concat_vec ('b"01" : mword 2) creg.
+
+Definition zreg : regidx := 'b"00000" : mword 5.
Hint Unfold zreg : sail.
-Definition ra : regidx := (vec_of_bits [B0;B0;B0;B0;B1] : mword 5).
+Definition ra : regidx := 'b"00001" : mword 5.
Hint Unfold ra : sail.
-Definition sp : regidx := (vec_of_bits [B0;B0;B0;B1;B0] : mword 5).
+Definition sp : regidx := 'b"00010" : mword 5.
Hint Unfold sp : sail.
-Definition Architecture_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
-: Architecture :=
-
- let l__166 := arg_ in
- if sumbool_of_bool ((Z.eqb l__166 0)) then RV32
- else if sumbool_of_bool ((Z.eqb l__166 1)) then RV64
+Definition Architecture_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : Architecture :=
+ let l__262 := arg_ in
+ if sumbool_of_bool (Z.eqb l__262 0) then RV32
+ else if sumbool_of_bool (Z.eqb l__262 1) then RV64
else RV128.
-Definition num_of_Architecture (arg_ : Architecture)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | RV32 => 0 | RV64 => 1 | RV128 => 2 end).
+Definition num_of_Architecture (arg_ : Architecture) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | RV32 => 0 | RV64 => 1 | RV128 => 2 end).
-Definition architecture (a : mword 2)
-: option Architecture :=
-
+Definition architecture (a : mword 2) : option Architecture :=
let b__0 := a in
- if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then Some (RV32)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then Some (RV64)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then Some (RV128)
+ if eq_vec b__0 ('b"01" : mword 2) then Some RV32
+ else if eq_vec b__0 ('b"10" : mword 2) then Some RV64
+ else if eq_vec b__0 ('b"11" : mword 2) then Some RV128
else None.
-Definition arch_to_bits (a : Architecture)
-: mword 2 :=
-
+Definition arch_to_bits (a : Architecture) : mword 2 :=
match a with
- | RV32 => (vec_of_bits [B0;B1] : mword 2)
- | RV64 => (vec_of_bits [B1;B0] : mword 2)
- | RV128 => (vec_of_bits [B1;B1] : mword 2)
+ | RV32 => 'b"01" : mword 2
+ | RV64 => 'b"10" : mword 2
+ | RV128 => 'b"11" : mword 2
end.
-Definition Privilege_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
-: Privilege :=
-
- let l__164 := arg_ in
- if sumbool_of_bool ((Z.eqb l__164 0)) then User
- else if sumbool_of_bool ((Z.eqb l__164 1)) then Supervisor
+Definition Privilege_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : Privilege :=
+ let l__260 := arg_ in
+ if sumbool_of_bool (Z.eqb l__260 0) then User
+ else if sumbool_of_bool (Z.eqb l__260 1) then Supervisor
else Machine.
-Definition num_of_Privilege (arg_ : Privilege)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | User => 0 | Supervisor => 1 | Machine => 2 end).
+Definition num_of_Privilege (arg_ : Privilege) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | User => 0 | Supervisor => 1 | Machine => 2 end).
-Definition privLevel_to_bits (p : Privilege)
-: mword 2 :=
-
+Definition privLevel_to_bits (p : Privilege) : mword 2 :=
match p with
- | User => (vec_of_bits [B0;B0] : mword 2)
- | Supervisor => (vec_of_bits [B0;B1] : mword 2)
- | Machine => (vec_of_bits [B1;B1] : mword 2)
+ | User => 'b"00" : mword 2
+ | Supervisor => 'b"01" : mword 2
+ | Machine => 'b"11" : mword 2
end.
-Definition privLevel_of_bits (p : mword 2)
-: M (Privilege) :=
-
+Definition privLevel_of_bits (p : mword 2) : M (Privilege) :=
let b__0 := p in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0] : mword 2))) then returnm (User : Privilege)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then returnm (Supervisor : Privilege)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then returnm (Machine : Privilege)
+ (if eq_vec b__0 ('b"00" : mword 2) then returnm User
+ else if eq_vec b__0 ('b"01" : mword 2) then returnm Supervisor
+ else if eq_vec b__0 ('b"11" : mword 2) then returnm Machine
else
assert_exp' false "Pattern match failure at model/riscv_types.sail 78:2 - 82:3" >>= fun _ =>
exit tt)
: M (Privilege).
-Definition privLevel_to_str (p : Privilege)
-: string :=
-
+Definition privLevel_to_str (p : Privilege) : string :=
match p with | User => "U" | Supervisor => "S" | Machine => "M" end.
-Definition Retired_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 1)}
-: Retired :=
-
- let l__163 := arg_ in
- if sumbool_of_bool ((Z.eqb l__163 0)) then RETIRE_SUCCESS
+Definition Retired_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 1))} : Retired :=
+ let l__259 := arg_ in
+ if sumbool_of_bool (Z.eqb l__259 0) then RETIRE_SUCCESS
else RETIRE_FAIL.
-Definition num_of_Retired (arg_ : Retired)
-: {e : Z & ArithFact (0 <= e /\ e <= 1)} :=
-
- build_ex(match arg_ with | RETIRE_SUCCESS => 0 | RETIRE_FAIL => 1 end).
-
-Definition AccessType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
-: AccessType :=
-
- let l__160 := arg_ in
- if sumbool_of_bool ((Z.eqb l__160 0)) then Read
- else if sumbool_of_bool ((Z.eqb l__160 1)) then Write
- else if sumbool_of_bool ((Z.eqb l__160 2)) then ReadWrite
- else Execute.
-
-Definition num_of_AccessType (arg_ : AccessType)
-: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
-
- build_ex(match arg_ with | Read => 0 | Write => 1 | ReadWrite => 2 | Execute => 3 end).
-
-Definition accessType_to_str (a : AccessType)
-: string :=
-
- match a with | Read => "R" | Write => "W" | ReadWrite => "RW" | Execute => "X" end.
-
-Definition word_width_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
-: word_width :=
-
- let l__157 := arg_ in
- if sumbool_of_bool ((Z.eqb l__157 0)) then BYTE
- else if sumbool_of_bool ((Z.eqb l__157 1)) then HALF
- else if sumbool_of_bool ((Z.eqb l__157 2)) then WORD
+Definition num_of_Retired (arg_ : Retired) : {e : Z & ArithFact ((0 <=? e) && (e <=? 1))} :=
+ build_ex (match arg_ with | RETIRE_SUCCESS => 0 | RETIRE_FAIL => 1 end).
+
+Definition word_width_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))} : word_width :=
+ let l__256 := arg_ in
+ if sumbool_of_bool (Z.eqb l__256 0) then BYTE
+ else if sumbool_of_bool (Z.eqb l__256 1) then HALF
+ else if sumbool_of_bool (Z.eqb l__256 2) then WORD
else DOUBLE.
-Definition num_of_word_width (arg_ : word_width)
-: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
-
- build_ex(match arg_ with | BYTE => 0 | HALF => 1 | WORD => 2 | DOUBLE => 3 end).
+Definition num_of_word_width (arg_ : word_width) : {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (match arg_ with | BYTE => 0 | HALF => 1 | WORD => 2 | DOUBLE => 3 end).
-Definition InterruptType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 8)}
+Definition InterruptType_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 8))}
: InterruptType :=
-
- let l__149 := arg_ in
- if sumbool_of_bool ((Z.eqb l__149 0)) then I_U_Software
- else if sumbool_of_bool ((Z.eqb l__149 1)) then I_S_Software
- else if sumbool_of_bool ((Z.eqb l__149 2)) then I_M_Software
- else if sumbool_of_bool ((Z.eqb l__149 3)) then I_U_Timer
- else if sumbool_of_bool ((Z.eqb l__149 4)) then I_S_Timer
- else if sumbool_of_bool ((Z.eqb l__149 5)) then I_M_Timer
- else if sumbool_of_bool ((Z.eqb l__149 6)) then I_U_External
- else if sumbool_of_bool ((Z.eqb l__149 7)) then I_S_External
+ let l__248 := arg_ in
+ if sumbool_of_bool (Z.eqb l__248 0) then I_U_Software
+ else if sumbool_of_bool (Z.eqb l__248 1) then I_S_Software
+ else if sumbool_of_bool (Z.eqb l__248 2) then I_M_Software
+ else if sumbool_of_bool (Z.eqb l__248 3) then I_U_Timer
+ else if sumbool_of_bool (Z.eqb l__248 4) then I_S_Timer
+ else if sumbool_of_bool (Z.eqb l__248 5) then I_M_Timer
+ else if sumbool_of_bool (Z.eqb l__248 6) then I_U_External
+ else if sumbool_of_bool (Z.eqb l__248 7) then I_S_External
else I_M_External.
-Definition num_of_InterruptType (arg_ : InterruptType)
-: {e : Z & ArithFact (0 <= e /\ e <= 8)} :=
-
- build_ex(match arg_ with
- | I_U_Software => 0
- | I_S_Software => 1
- | I_M_Software => 2
- | I_U_Timer => 3
- | I_S_Timer => 4
- | I_M_Timer => 5
- | I_U_External => 6
- | I_S_External => 7
- | I_M_External => 8
- end).
-
-Definition interruptType_to_bits (i : InterruptType)
-: mword 8 :=
-
+Definition num_of_InterruptType (arg_ : InterruptType)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 8))} :=
+ build_ex (
+ match arg_ with
+ | I_U_Software => 0
+ | I_S_Software => 1
+ | I_M_Software => 2
+ | I_U_Timer => 3
+ | I_S_Timer => 4
+ | I_M_Timer => 5
+ | I_U_External => 6
+ | I_S_External => 7
+ | I_M_External => 8
+ end
+ ).
+
+Definition interruptType_to_bits (i : InterruptType) : mword 8 :=
match i with
- | I_U_Software => (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8)
- | I_S_Software => (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B1] : mword 8)
- | I_M_Software => (vec_of_bits [B0;B0;B0;B0;B0;B0;B1;B1] : mword 8)
- | I_U_Timer => (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0] : mword 8)
- | I_S_Timer => (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B1] : mword 8)
- | I_M_Timer => (vec_of_bits [B0;B0;B0;B0;B0;B1;B1;B1] : mword 8)
- | I_U_External => (vec_of_bits [B0;B0;B0;B0;B1;B0;B0;B0] : mword 8)
- | I_S_External => (vec_of_bits [B0;B0;B0;B0;B1;B0;B0;B1] : mword 8)
- | I_M_External => (vec_of_bits [B0;B0;B0;B0;B1;B0;B1;B1] : mword 8)
+ | I_U_Software => Ox"00" : mword 8
+ | I_S_Software => Ox"01" : mword 8
+ | I_M_Software => Ox"03" : mword 8
+ | I_U_Timer => Ox"04" : mword 8
+ | I_S_Timer => Ox"05" : mword 8
+ | I_M_Timer => Ox"07" : mword 8
+ | I_U_External => Ox"08" : mword 8
+ | I_S_External => Ox"09" : mword 8
+ | I_M_External => Ox"0B" : mword 8
end.
-Definition ExceptionType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 16)}
-: ExceptionType :=
-
- let l__133 := arg_ in
- if sumbool_of_bool ((Z.eqb l__133 0)) then E_Fetch_Addr_Align
- else if sumbool_of_bool ((Z.eqb l__133 1)) then E_Fetch_Access_Fault
- else if sumbool_of_bool ((Z.eqb l__133 2)) then E_Illegal_Instr
- else if sumbool_of_bool ((Z.eqb l__133 3)) then E_Breakpoint
- else if sumbool_of_bool ((Z.eqb l__133 4)) then E_Load_Addr_Align
- else if sumbool_of_bool ((Z.eqb l__133 5)) then E_Load_Access_Fault
- else if sumbool_of_bool ((Z.eqb l__133 6)) then E_SAMO_Addr_Align
- else if sumbool_of_bool ((Z.eqb l__133 7)) then E_SAMO_Access_Fault
- else if sumbool_of_bool ((Z.eqb l__133 8)) then E_U_EnvCall
- else if sumbool_of_bool ((Z.eqb l__133 9)) then E_S_EnvCall
- else if sumbool_of_bool ((Z.eqb l__133 10)) then E_Reserved_10
- else if sumbool_of_bool ((Z.eqb l__133 11)) then E_M_EnvCall
- else if sumbool_of_bool ((Z.eqb l__133 12)) then E_Fetch_Page_Fault
- else if sumbool_of_bool ((Z.eqb l__133 13)) then E_Load_Page_Fault
- else if sumbool_of_bool ((Z.eqb l__133 14)) then E_Reserved_14
- else if sumbool_of_bool ((Z.eqb l__133 15)) then E_SAMO_Page_Fault
- else E_CHERI.
-
-Definition num_of_ExceptionType (arg_ : ExceptionType)
-: {e : Z & ArithFact (0 <= e /\ e <= 16)} :=
-
- build_ex(match arg_ with
- | E_Fetch_Addr_Align => 0
- | E_Fetch_Access_Fault => 1
- | E_Illegal_Instr => 2
- | E_Breakpoint => 3
- | E_Load_Addr_Align => 4
- | E_Load_Access_Fault => 5
- | E_SAMO_Addr_Align => 6
- | E_SAMO_Access_Fault => 7
- | E_U_EnvCall => 8
- | E_S_EnvCall => 9
- | E_Reserved_10 => 10
- | E_M_EnvCall => 11
- | E_Fetch_Page_Fault => 12
- | E_Load_Page_Fault => 13
- | E_Reserved_14 => 14
- | E_SAMO_Page_Fault => 15
- | E_CHERI => 16
- end).
-
-Definition exceptionType_to_bits (e : ExceptionType)
-: mword 8 :=
-
+Definition exceptionType_to_bits (e : ExceptionType) : mword 8 :=
match e with
- | E_Fetch_Addr_Align => (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8)
- | E_Fetch_Access_Fault => (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B1] : mword 8)
- | E_Illegal_Instr => (vec_of_bits [B0;B0;B0;B0;B0;B0;B1;B0] : mword 8)
- | E_Breakpoint => (vec_of_bits [B0;B0;B0;B0;B0;B0;B1;B1] : mword 8)
- | E_Load_Addr_Align => (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0] : mword 8)
- | E_Load_Access_Fault => (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B1] : mword 8)
- | E_SAMO_Addr_Align => (vec_of_bits [B0;B0;B0;B0;B0;B1;B1;B0] : mword 8)
- | E_SAMO_Access_Fault => (vec_of_bits [B0;B0;B0;B0;B0;B1;B1;B1] : mword 8)
- | E_U_EnvCall => (vec_of_bits [B0;B0;B0;B0;B1;B0;B0;B0] : mword 8)
- | E_S_EnvCall => (vec_of_bits [B0;B0;B0;B0;B1;B0;B0;B1] : mword 8)
- | E_Reserved_10 => (vec_of_bits [B0;B0;B0;B0;B1;B0;B1;B0] : mword 8)
- | E_M_EnvCall => (vec_of_bits [B0;B0;B0;B0;B1;B0;B1;B1] : mword 8)
- | E_Fetch_Page_Fault => (vec_of_bits [B0;B0;B0;B0;B1;B1;B0;B0] : mword 8)
- | E_Load_Page_Fault => (vec_of_bits [B0;B0;B0;B0;B1;B1;B0;B1] : mword 8)
- | E_Reserved_14 => (vec_of_bits [B0;B0;B0;B0;B1;B1;B1;B0] : mword 8)
- | E_SAMO_Page_Fault => (vec_of_bits [B0;B0;B0;B0;B1;B1;B1;B1] : mword 8)
- | E_CHERI => (vec_of_bits [B0;B0;B1;B0;B0;B0;B0;B0] : mword 8)
+ | E_Fetch_Addr_Align tt => Ox"00" : mword 8
+ | E_Fetch_Access_Fault tt => Ox"01" : mword 8
+ | E_Illegal_Instr tt => Ox"02" : mword 8
+ | E_Breakpoint tt => Ox"03" : mword 8
+ | E_Load_Addr_Align tt => Ox"04" : mword 8
+ | E_Load_Access_Fault tt => Ox"05" : mword 8
+ | E_SAMO_Addr_Align tt => Ox"06" : mword 8
+ | E_SAMO_Access_Fault tt => Ox"07" : mword 8
+ | E_U_EnvCall tt => Ox"08" : mword 8
+ | E_S_EnvCall tt => Ox"09" : mword 8
+ | E_Reserved_10 tt => Ox"0A" : mword 8
+ | E_M_EnvCall tt => Ox"0B" : mword 8
+ | E_Fetch_Page_Fault tt => Ox"0C" : mword 8
+ | E_Load_Page_Fault tt => Ox"0D" : mword 8
+ | E_Reserved_14 tt => Ox"0E" : mword 8
+ | E_SAMO_Page_Fault tt => Ox"0F" : mword 8
+ | E_Extension e => ext_exc_type_to_bits e
end.
-Definition exceptionType_to_str (e : ExceptionType)
-: string :=
-
+Definition num_of_ExceptionType (e : ExceptionType) : {n : Z & ArithFact ((0 <=? n) && (n <? xlen))} :=
+ build_ex (
+ match e with
+ | E_Fetch_Addr_Align tt => 0
+ | E_Fetch_Access_Fault tt => 1
+ | E_Illegal_Instr tt => 2
+ | E_Breakpoint tt => 3
+ | E_Load_Addr_Align tt => 4
+ | E_Load_Access_Fault tt => 5
+ | E_SAMO_Addr_Align tt => 6
+ | E_SAMO_Access_Fault tt => 7
+ | E_U_EnvCall tt => 8
+ | E_S_EnvCall tt => 9
+ | E_Reserved_10 tt => 10
+ | E_M_EnvCall tt => 11
+ | E_Fetch_Page_Fault tt => 12
+ | E_Load_Page_Fault tt => 13
+ | E_Reserved_14 tt => 14
+ | E_SAMO_Page_Fault tt => 15
+ | E_Extension e => projT1 (num_of_ext_exc_type e)
+ end
+ ).
+
+Definition exceptionType_to_str (e : ExceptionType) : string :=
match e with
- | E_Fetch_Addr_Align => "misaligned-fetch"
- | E_Fetch_Access_Fault => "fetch-access-fault"
- | E_Illegal_Instr => "illegal-instruction"
- | E_Breakpoint => "breakpoint"
- | E_Load_Addr_Align => "misaligned-load"
- | E_Load_Access_Fault => "load-access-fault"
- | E_SAMO_Addr_Align => "misaliged-store/amo"
- | E_SAMO_Access_Fault => "store/amo-access-fault"
- | E_U_EnvCall => "u-call"
- | E_S_EnvCall => "s-call"
- | E_Reserved_10 => "reserved-0"
- | E_M_EnvCall => "m-call"
- | E_Fetch_Page_Fault => "fetch-page-fault"
- | E_Load_Page_Fault => "load-page-fault"
- | E_Reserved_14 => "reserved-1"
- | E_SAMO_Page_Fault => "store/amo-page-fault"
- | E_CHERI => "CHERI"
+ | E_Fetch_Addr_Align tt => "misaligned-fetch"
+ | E_Fetch_Access_Fault tt => "fetch-access-fault"
+ | E_Illegal_Instr tt => "illegal-instruction"
+ | E_Breakpoint tt => "breakpoint"
+ | E_Load_Addr_Align tt => "misaligned-load"
+ | E_Load_Access_Fault tt => "load-access-fault"
+ | E_SAMO_Addr_Align tt => "misaliged-store/amo"
+ | E_SAMO_Access_Fault tt => "store/amo-access-fault"
+ | E_U_EnvCall tt => "u-call"
+ | E_S_EnvCall tt => "s-call"
+ | E_Reserved_10 tt => "reserved-0"
+ | E_M_EnvCall tt => "m-call"
+ | E_Fetch_Page_Fault tt => "fetch-page-fault"
+ | E_Load_Page_Fault tt => "load-page-fault"
+ | E_Reserved_14 tt => "reserved-1"
+ | E_SAMO_Page_Fault tt => "store/amo-page-fault"
+ | E_Extension e => ext_exc_type_to_str e
end.
-Definition not_implemented {a : Type} (message : string)
-: M (a) :=
-
- throw (Error_not_implemented
- (message)).
+Definition not_implemented {a : Type} (message : string) : M (a) :=
+ throw (Error_not_implemented message).
-Definition internal_error {a : Type} (s : string)
-: M (a) :=
-
+Definition internal_error {a : Type} (s : string) : M (a) :=
assert_exp' false s >>= fun _ => exit tt.
-Definition TrapVectorMode_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
+Definition TrapVectorMode_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))}
: TrapVectorMode :=
-
- let l__131 := arg_ in
- if sumbool_of_bool ((Z.eqb l__131 0)) then TV_Direct
- else if sumbool_of_bool ((Z.eqb l__131 1)) then TV_Vector
+ let l__246 := arg_ in
+ if sumbool_of_bool (Z.eqb l__246 0) then TV_Direct
+ else if sumbool_of_bool (Z.eqb l__246 1) then TV_Vector
else TV_Reserved.
-Definition num_of_TrapVectorMode (arg_ : TrapVectorMode)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | TV_Direct => 0 | TV_Vector => 1 | TV_Reserved => 2 end).
+Definition num_of_TrapVectorMode (arg_ : TrapVectorMode)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | TV_Direct => 0 | TV_Vector => 1 | TV_Reserved => 2 end).
-Definition trapVectorMode_of_bits (m : mword 2)
-: TrapVectorMode :=
-
+Definition trapVectorMode_of_bits (m : mword 2) : TrapVectorMode :=
let b__0 := m in
- if ((eq_vec b__0 (vec_of_bits [B0;B0] : mword 2))) then TV_Direct
- else if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then TV_Vector
+ if eq_vec b__0 ('b"00" : mword 2) then TV_Direct
+ else if eq_vec b__0 ('b"01" : mword 2) then TV_Vector
else TV_Reserved.
-Definition ExtStatus_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
-: ExtStatus :=
-
- let l__128 := arg_ in
- if sumbool_of_bool ((Z.eqb l__128 0)) then Off
- else if sumbool_of_bool ((Z.eqb l__128 1)) then Initial
- else if sumbool_of_bool ((Z.eqb l__128 2)) then Clean
+Definition ExtStatus_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))} : ExtStatus :=
+ let l__243 := arg_ in
+ if sumbool_of_bool (Z.eqb l__243 0) then Off
+ else if sumbool_of_bool (Z.eqb l__243 1) then Initial
+ else if sumbool_of_bool (Z.eqb l__243 2) then Clean
else Dirty.
-Definition num_of_ExtStatus (arg_ : ExtStatus)
-: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
-
- build_ex(match arg_ with | Off => 0 | Initial => 1 | Clean => 2 | Dirty => 3 end).
+Definition num_of_ExtStatus (arg_ : ExtStatus) : {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (match arg_ with | Off => 0 | Initial => 1 | Clean => 2 | Dirty => 3 end).
-Definition extStatus_to_bits (e : ExtStatus)
-: mword 2 :=
-
+Definition extStatus_to_bits (e : ExtStatus) : mword 2 :=
match e with
- | Off => (vec_of_bits [B0;B0] : mword 2)
- | Initial => (vec_of_bits [B0;B1] : mword 2)
- | Clean => (vec_of_bits [B1;B0] : mword 2)
- | Dirty => (vec_of_bits [B1;B1] : mword 2)
+ | Off => 'b"00" : mword 2
+ | Initial => 'b"01" : mword 2
+ | Clean => 'b"10" : mword 2
+ | Dirty => 'b"11" : mword 2
end.
-Definition extStatus_of_bits (e : mword 2)
-: M (ExtStatus) :=
-
+Definition extStatus_of_bits (e : mword 2) : M (ExtStatus) :=
let b__0 := e in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0] : mword 2))) then returnm (Off : ExtStatus)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then returnm (Initial : ExtStatus)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then returnm (Clean : ExtStatus)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then returnm (Dirty : ExtStatus)
+ (if eq_vec b__0 ('b"00" : mword 2) then returnm Off
+ else if eq_vec b__0 ('b"01" : mword 2) then returnm Initial
+ else if eq_vec b__0 ('b"10" : mword 2) then returnm Clean
+ else if eq_vec b__0 ('b"11" : mword 2) then returnm Dirty
else
- assert_exp' false "Pattern match failure at model/riscv_types.sail 264:2 - 269:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/riscv_types.sail 281:2 - 286:3" >>= fun _ =>
exit tt)
: M (ExtStatus).
-Definition SATPMode_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
-: SATPMode :=
-
- let l__125 := arg_ in
- if sumbool_of_bool ((Z.eqb l__125 0)) then Sbare
- else if sumbool_of_bool ((Z.eqb l__125 1)) then Sv32
- else if sumbool_of_bool ((Z.eqb l__125 2)) then Sv39
+Definition SATPMode_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))} : SATPMode :=
+ let l__240 := arg_ in
+ if sumbool_of_bool (Z.eqb l__240 0) then Sbare
+ else if sumbool_of_bool (Z.eqb l__240 1) then Sv32
+ else if sumbool_of_bool (Z.eqb l__240 2) then Sv39
else Sv48.
-Definition num_of_SATPMode (arg_ : SATPMode)
-: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
-
- build_ex(match arg_ with | Sbare => 0 | Sv32 => 1 | Sv39 => 2 | Sv48 => 3 end).
+Definition num_of_SATPMode (arg_ : SATPMode) : {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (match arg_ with | Sbare => 0 | Sv32 => 1 | Sv39 => 2 | Sv48 => 3 end).
-Definition satp64Mode_of_bits (a : Architecture) (m : mword 4)
-: option SATPMode :=
-
+Definition satp64Mode_of_bits (a : Architecture) (m : mword 4) : option SATPMode :=
match (a, m) with
- | (g__223, b__0) =>
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0] : mword 4))) then Some (Sbare)
+ | (g__326, b__0) =>
+ if eq_vec b__0 (Ox"0" : mword 4) then Some Sbare
else
- match (g__223, b__0) with
+ match (g__326, b__0) with
| (RV32, b__0) =>
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1] : mword 4))) then Some (Sv32)
+ if eq_vec b__0 (Ox"1" : mword 4) then Some Sv32
else match (RV32, b__0) with | (_, _) => None end
| (RV64, b__0) =>
- if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0] : mword 4))) then Some (Sv39)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B1] : mword 4))) then Some (Sv48)
+ if eq_vec b__0 (Ox"8" : mword 4) then Some Sv39
+ else if eq_vec b__0 (Ox"9" : mword 4) then Some Sv48
else match (RV64, b__0) with | (_, _) => None end
| (_, _) => None
end
end.
-Definition uop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 1)}
-: uop :=
-
- let l__124 := arg_ in
- if sumbool_of_bool ((Z.eqb l__124 0)) then RISCV_LUI
+Definition uop_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 1))} : uop :=
+ let l__239 := arg_ in
+ if sumbool_of_bool (Z.eqb l__239 0) then RISCV_LUI
else RISCV_AUIPC.
-Definition num_of_uop (arg_ : uop)
-: {e : Z & ArithFact (0 <= e /\ e <= 1)} :=
-
- build_ex(match arg_ with | RISCV_LUI => 0 | RISCV_AUIPC => 1 end).
-
-Definition bop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 5)}
-: bop :=
-
- let l__119 := arg_ in
- if sumbool_of_bool ((Z.eqb l__119 0)) then RISCV_BEQ
- else if sumbool_of_bool ((Z.eqb l__119 1)) then RISCV_BNE
- else if sumbool_of_bool ((Z.eqb l__119 2)) then RISCV_BLT
- else if sumbool_of_bool ((Z.eqb l__119 3)) then RISCV_BGE
- else if sumbool_of_bool ((Z.eqb l__119 4)) then RISCV_BLTU
- else RISCV_BGEU.
+Definition num_of_uop (arg_ : uop) : {e : Z & ArithFact ((0 <=? e) && (e <=? 1))} :=
+ build_ex (match arg_ with | RISCV_LUI => 0 | RISCV_AUIPC => 1 end).
-Definition num_of_bop (arg_ : bop)
-: {e : Z & ArithFact (0 <= e /\ e <= 5)} :=
-
- build_ex(match arg_ with
- | RISCV_BEQ => 0
- | RISCV_BNE => 1
- | RISCV_BLT => 2
- | RISCV_BGE => 3
- | RISCV_BLTU => 4
- | RISCV_BGEU => 5
- end).
+Definition bop_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 5))} : bop :=
+ let l__234 := arg_ in
+ if sumbool_of_bool (Z.eqb l__234 0) then RISCV_BEQ
+ else if sumbool_of_bool (Z.eqb l__234 1) then RISCV_BNE
+ else if sumbool_of_bool (Z.eqb l__234 2) then RISCV_BLT
+ else if sumbool_of_bool (Z.eqb l__234 3) then RISCV_BGE
+ else if sumbool_of_bool (Z.eqb l__234 4) then RISCV_BLTU
+ else RISCV_BGEU.
-Definition iop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 5)}
-: iop :=
-
- let l__114 := arg_ in
- if sumbool_of_bool ((Z.eqb l__114 0)) then RISCV_ADDI
- else if sumbool_of_bool ((Z.eqb l__114 1)) then RISCV_SLTI
- else if sumbool_of_bool ((Z.eqb l__114 2)) then RISCV_SLTIU
- else if sumbool_of_bool ((Z.eqb l__114 3)) then RISCV_XORI
- else if sumbool_of_bool ((Z.eqb l__114 4)) then RISCV_ORI
+Definition num_of_bop (arg_ : bop) : {e : Z & ArithFact ((0 <=? e) && (e <=? 5))} :=
+ build_ex (
+ match arg_ with
+ | RISCV_BEQ => 0
+ | RISCV_BNE => 1
+ | RISCV_BLT => 2
+ | RISCV_BGE => 3
+ | RISCV_BLTU => 4
+ | RISCV_BGEU => 5
+ end
+ ).
+
+Definition iop_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 5))} : iop :=
+ let l__229 := arg_ in
+ if sumbool_of_bool (Z.eqb l__229 0) then RISCV_ADDI
+ else if sumbool_of_bool (Z.eqb l__229 1) then RISCV_SLTI
+ else if sumbool_of_bool (Z.eqb l__229 2) then RISCV_SLTIU
+ else if sumbool_of_bool (Z.eqb l__229 3) then RISCV_XORI
+ else if sumbool_of_bool (Z.eqb l__229 4) then RISCV_ORI
else RISCV_ANDI.
-Definition num_of_iop (arg_ : iop)
-: {e : Z & ArithFact (0 <= e /\ e <= 5)} :=
-
- build_ex(match arg_ with
- | RISCV_ADDI => 0
- | RISCV_SLTI => 1
- | RISCV_SLTIU => 2
- | RISCV_XORI => 3
- | RISCV_ORI => 4
- | RISCV_ANDI => 5
- end).
-
-Definition sop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
-: sop :=
-
- let l__112 := arg_ in
- if sumbool_of_bool ((Z.eqb l__112 0)) then RISCV_SLLI
- else if sumbool_of_bool ((Z.eqb l__112 1)) then RISCV_SRLI
+Definition num_of_iop (arg_ : iop) : {e : Z & ArithFact ((0 <=? e) && (e <=? 5))} :=
+ build_ex (
+ match arg_ with
+ | RISCV_ADDI => 0
+ | RISCV_SLTI => 1
+ | RISCV_SLTIU => 2
+ | RISCV_XORI => 3
+ | RISCV_ORI => 4
+ | RISCV_ANDI => 5
+ end
+ ).
+
+Definition sop_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : sop :=
+ let l__227 := arg_ in
+ if sumbool_of_bool (Z.eqb l__227 0) then RISCV_SLLI
+ else if sumbool_of_bool (Z.eqb l__227 1) then RISCV_SRLI
else RISCV_SRAI.
-Definition num_of_sop (arg_ : sop)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | RISCV_SLLI => 0 | RISCV_SRLI => 1 | RISCV_SRAI => 2 end).
-
-Definition rop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 9)}
-: rop :=
-
- let l__103 := arg_ in
- if sumbool_of_bool ((Z.eqb l__103 0)) then RISCV_ADD
- else if sumbool_of_bool ((Z.eqb l__103 1)) then RISCV_SUB
- else if sumbool_of_bool ((Z.eqb l__103 2)) then RISCV_SLL
- else if sumbool_of_bool ((Z.eqb l__103 3)) then RISCV_SLT
- else if sumbool_of_bool ((Z.eqb l__103 4)) then RISCV_SLTU
- else if sumbool_of_bool ((Z.eqb l__103 5)) then RISCV_XOR
- else if sumbool_of_bool ((Z.eqb l__103 6)) then RISCV_SRL
- else if sumbool_of_bool ((Z.eqb l__103 7)) then RISCV_SRA
- else if sumbool_of_bool ((Z.eqb l__103 8)) then RISCV_OR
+Definition num_of_sop (arg_ : sop) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | RISCV_SLLI => 0 | RISCV_SRLI => 1 | RISCV_SRAI => 2 end).
+
+Definition rop_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 9))} : rop :=
+ let l__218 := arg_ in
+ if sumbool_of_bool (Z.eqb l__218 0) then RISCV_ADD
+ else if sumbool_of_bool (Z.eqb l__218 1) then RISCV_SUB
+ else if sumbool_of_bool (Z.eqb l__218 2) then RISCV_SLL
+ else if sumbool_of_bool (Z.eqb l__218 3) then RISCV_SLT
+ else if sumbool_of_bool (Z.eqb l__218 4) then RISCV_SLTU
+ else if sumbool_of_bool (Z.eqb l__218 5) then RISCV_XOR
+ else if sumbool_of_bool (Z.eqb l__218 6) then RISCV_SRL
+ else if sumbool_of_bool (Z.eqb l__218 7) then RISCV_SRA
+ else if sumbool_of_bool (Z.eqb l__218 8) then RISCV_OR
else RISCV_AND.
-Definition num_of_rop (arg_ : rop)
-: {e : Z & ArithFact (0 <= e /\ e <= 9)} :=
-
- build_ex(match arg_ with
- | RISCV_ADD => 0
- | RISCV_SUB => 1
- | RISCV_SLL => 2
- | RISCV_SLT => 3
- | RISCV_SLTU => 4
- | RISCV_XOR => 5
- | RISCV_SRL => 6
- | RISCV_SRA => 7
- | RISCV_OR => 8
- | RISCV_AND => 9
- end).
-
-Definition ropw_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 4)}
-: ropw :=
-
- let l__99 := arg_ in
- if sumbool_of_bool ((Z.eqb l__99 0)) then RISCV_ADDW
- else if sumbool_of_bool ((Z.eqb l__99 1)) then RISCV_SUBW
- else if sumbool_of_bool ((Z.eqb l__99 2)) then RISCV_SLLW
- else if sumbool_of_bool ((Z.eqb l__99 3)) then RISCV_SRLW
+Definition num_of_rop (arg_ : rop) : {e : Z & ArithFact ((0 <=? e) && (e <=? 9))} :=
+ build_ex (
+ match arg_ with
+ | RISCV_ADD => 0
+ | RISCV_SUB => 1
+ | RISCV_SLL => 2
+ | RISCV_SLT => 3
+ | RISCV_SLTU => 4
+ | RISCV_XOR => 5
+ | RISCV_SRL => 6
+ | RISCV_SRA => 7
+ | RISCV_OR => 8
+ | RISCV_AND => 9
+ end
+ ).
+
+Definition ropw_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 4))} : ropw :=
+ let l__214 := arg_ in
+ if sumbool_of_bool (Z.eqb l__214 0) then RISCV_ADDW
+ else if sumbool_of_bool (Z.eqb l__214 1) then RISCV_SUBW
+ else if sumbool_of_bool (Z.eqb l__214 2) then RISCV_SLLW
+ else if sumbool_of_bool (Z.eqb l__214 3) then RISCV_SRLW
else RISCV_SRAW.
-Definition num_of_ropw (arg_ : ropw)
-: {e : Z & ArithFact (0 <= e /\ e <= 4)} :=
-
- build_ex(match arg_ with
- | RISCV_ADDW => 0
- | RISCV_SUBW => 1
- | RISCV_SLLW => 2
- | RISCV_SRLW => 3
- | RISCV_SRAW => 4
- end).
-
-Definition sopw_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
-: sopw :=
-
- let l__97 := arg_ in
- if sumbool_of_bool ((Z.eqb l__97 0)) then RISCV_SLLIW
- else if sumbool_of_bool ((Z.eqb l__97 1)) then RISCV_SRLIW
+Definition num_of_ropw (arg_ : ropw) : {e : Z & ArithFact ((0 <=? e) && (e <=? 4))} :=
+ build_ex (
+ match arg_ with
+ | RISCV_ADDW => 0
+ | RISCV_SUBW => 1
+ | RISCV_SLLW => 2
+ | RISCV_SRLW => 3
+ | RISCV_SRAW => 4
+ end
+ ).
+
+Definition sopw_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : sopw :=
+ let l__212 := arg_ in
+ if sumbool_of_bool (Z.eqb l__212 0) then RISCV_SLLIW
+ else if sumbool_of_bool (Z.eqb l__212 1) then RISCV_SRLIW
else RISCV_SRAIW.
-Definition num_of_sopw (arg_ : sopw)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | RISCV_SLLIW => 0 | RISCV_SRLIW => 1 | RISCV_SRAIW => 2 end).
-
-Definition amoop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 8)}
-: amoop :=
-
- let l__89 := arg_ in
- if sumbool_of_bool ((Z.eqb l__89 0)) then AMOSWAP
- else if sumbool_of_bool ((Z.eqb l__89 1)) then AMOADD
- else if sumbool_of_bool ((Z.eqb l__89 2)) then AMOXOR
- else if sumbool_of_bool ((Z.eqb l__89 3)) then AMOAND
- else if sumbool_of_bool ((Z.eqb l__89 4)) then AMOOR
- else if sumbool_of_bool ((Z.eqb l__89 5)) then AMOMIN
- else if sumbool_of_bool ((Z.eqb l__89 6)) then AMOMAX
- else if sumbool_of_bool ((Z.eqb l__89 7)) then AMOMINU
+Definition num_of_sopw (arg_ : sopw) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | RISCV_SLLIW => 0 | RISCV_SRLIW => 1 | RISCV_SRAIW => 2 end).
+
+Definition amoop_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 8))} : amoop :=
+ let l__204 := arg_ in
+ if sumbool_of_bool (Z.eqb l__204 0) then AMOSWAP
+ else if sumbool_of_bool (Z.eqb l__204 1) then AMOADD
+ else if sumbool_of_bool (Z.eqb l__204 2) then AMOXOR
+ else if sumbool_of_bool (Z.eqb l__204 3) then AMOAND
+ else if sumbool_of_bool (Z.eqb l__204 4) then AMOOR
+ else if sumbool_of_bool (Z.eqb l__204 5) then AMOMIN
+ else if sumbool_of_bool (Z.eqb l__204 6) then AMOMAX
+ else if sumbool_of_bool (Z.eqb l__204 7) then AMOMINU
else AMOMAXU.
-Definition num_of_amoop (arg_ : amoop)
-: {e : Z & ArithFact (0 <= e /\ e <= 8)} :=
-
- build_ex(match arg_ with
- | AMOSWAP => 0
- | AMOADD => 1
- | AMOXOR => 2
- | AMOAND => 3
- | AMOOR => 4
- | AMOMIN => 5
- | AMOMAX => 6
- | AMOMINU => 7
- | AMOMAXU => 8
- end).
-
-Definition csrop_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
-: csrop :=
-
- let l__87 := arg_ in
- if sumbool_of_bool ((Z.eqb l__87 0)) then CSRRW
- else if sumbool_of_bool ((Z.eqb l__87 1)) then CSRRS
+Definition num_of_amoop (arg_ : amoop) : {e : Z & ArithFact ((0 <=? e) && (e <=? 8))} :=
+ build_ex (
+ match arg_ with
+ | AMOSWAP => 0
+ | AMOADD => 1
+ | AMOXOR => 2
+ | AMOAND => 3
+ | AMOOR => 4
+ | AMOMIN => 5
+ | AMOMAX => 6
+ | AMOMINU => 7
+ | AMOMAXU => 8
+ end
+ ).
+
+Definition csrop_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : csrop :=
+ let l__202 := arg_ in
+ if sumbool_of_bool (Z.eqb l__202 0) then CSRRW
+ else if sumbool_of_bool (Z.eqb l__202 1) then CSRRS
else CSRRC.
-Definition num_of_csrop (arg_ : csrop)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | CSRRW => 0 | CSRRS => 1 | CSRRC => 2 end).
+Definition num_of_csrop (arg_ : csrop) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | CSRRW => 0 | CSRRS => 1 | CSRRC => 2 end).
-Definition sep_forwards (arg_ : unit)
-: string :=
-
+Definition sep_forwards (arg_ : unit) : string :=
match arg_ with
| tt =>
string_append (opt_spc_forwards tt)
(string_append "," (string_append (def_spc_forwards tt) ""))
end.
-Definition _s0_ (_s1_ : string)
-: M (option unit) :=
-
+Definition _s0_ (_s1_ : string) : M (option unit) :=
(match _s1_ with
| _s2_ =>
- (opt_spc_matches_prefix _s2_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ (opt_spc_matches_prefix _s2_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3_ _)) =>
+ | Some (tt, existT _ _s3_ _) =>
let _s4_ := string_drop _s2_ _s3_ in
- (if ((string_startswith _s4_ ",")) then
+ (if string_startswith _s4_ "," then
(match (string_drop _s4_ (projT1 (string_length ","))) with
| _s5_ =>
- (def_spc_matches_prefix _s5_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ (def_spc_matches_prefix _s5_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((tt, existT _ _s6_ _)) =>
- let p0_ := string_drop _s5_ _s6_ in
- if ((generic_eq p0_ "")) then Some (tt)
- else None
- | _ => None
- end)
- : option unit)
+ returnm (match w__1 with
+ | Some (tt, existT _ _s6_ _) =>
+ let p0_ := string_drop _s5_ _s6_ in
+ if generic_eq p0_ "" then Some tt
+ else None
+ | _ => None
+ end)
end)
: M (option unit)
- else returnm (None : option unit))
+ else returnm None)
: M (option unit)
- | _ => returnm (None : option unit)
+ | _ => returnm None
end)
: M (option unit)
end)
: M (option unit).
-Definition sep_backwards (arg_ : string)
-: M (unit) :=
-
+Definition sep_backwards (arg_ : string) : M (unit) :=
let _s7_ := arg_ in
(_s0_ _s7_) >>= fun w__0 : option unit =>
- (if ((match w__0 with | Some (tt) => true | _ => false end)) then
+ (if match w__0 with | Some tt => true | _ => false end then
(_s0_ _s7_) >>= fun w__1 : option unit =>
- (match w__1 with | Some (tt) => returnm (tt : unit) | _ => exit tt : M (unit) end)
+ (match w__1 with | Some tt => returnm tt | _ => exit tt : M (unit) end)
: M (unit)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (unit).
-Definition sep_forwards_matches (arg_ : unit) : bool := match arg_ with | tt => true end.
+Definition sep_forwards_matches (arg_ : unit) : bool := match arg_ with | tt => true end.
-Definition _s8_ (_s9_ : string)
-: M (option unit) :=
-
+Definition _s8_ (_s9_ : string) : M (option unit) :=
(match _s9_ with
| _s10_ =>
- (opt_spc_matches_prefix _s10_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ (opt_spc_matches_prefix _s10_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s11_ _)) =>
+ | Some (tt, existT _ _s11_ _) =>
let _s12_ := string_drop _s10_ _s11_ in
- (if ((string_startswith _s12_ ",")) then
+ (if string_startswith _s12_ "," then
(match (string_drop _s12_ (projT1 (string_length ","))) with
| _s13_ =>
- (def_spc_matches_prefix _s13_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ (def_spc_matches_prefix _s13_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((tt, existT _ _s14_ _)) =>
- let p0_ := string_drop _s13_ _s14_ in
- if ((generic_eq p0_ "")) then Some (tt)
- else None
- | _ => None
- end)
- : option unit)
+ returnm (match w__1 with
+ | Some (tt, existT _ _s14_ _) =>
+ let p0_ := string_drop _s13_ _s14_ in
+ if generic_eq p0_ "" then Some tt
+ else None
+ | _ => None
+ end)
end)
: M (option unit)
- else returnm (None : option unit))
+ else returnm None)
: M (option unit)
- | _ => returnm (None : option unit)
+ | _ => returnm None
end)
: M (option unit)
end)
: M (option unit).
-Definition sep_backwards_matches (arg_ : string)
-: M (bool) :=
-
+Definition sep_backwards_matches (arg_ : string) : M (bool) :=
let _s15_ := arg_ in
(_s8_ _s15_) >>= fun w__0 : option unit =>
- (if ((match w__0 with | Some (tt) => true | _ => false end)) then
+ (if match w__0 with | Some tt => true | _ => false end then
(_s8_ _s15_) >>= fun w__1 : option unit =>
- (match w__1 with
- | Some (tt) => returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | _ => exit tt : M (bool)
- end)
+ (match w__1 with | Some tt => returnm true | _ => exit tt : M (bool) end)
: M (bool)
- else returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))})))
+ else returnm false)
: M (bool).
-Definition _s16_ (_s17_ : string)
-: M (option string) :=
-
+Definition _s16_ (_s17_ : string) : M (option string) :=
(match _s17_ with
| _s18_ =>
- (opt_spc_matches_prefix _s18_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ (opt_spc_matches_prefix _s18_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s19_ _)) =>
+ | Some (tt, existT _ _s19_ _) =>
let _s20_ := string_drop _s18_ _s19_ in
- (if ((string_startswith _s20_ ",")) then
+ (if string_startswith _s20_ "," then
(match (string_drop _s20_ (projT1 (string_length ","))) with
| _s21_ =>
- (def_spc_matches_prefix _s21_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ (def_spc_matches_prefix _s21_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((tt, existT _ _s22_ _)) =>
- match (string_drop _s21_ _s22_) with | s_ => Some (s_) end
- | _ => None
- end)
- : option string)
+ returnm (match w__1 with
+ | Some (tt, existT _ _s22_ _) =>
+ match (string_drop _s21_ _s22_) with | s_ => Some s_ end
+ | _ => None
+ end)
end)
: M (option string)
- else returnm (None : option string))
+ else returnm None)
: M (option string)
- | _ => returnm (None : option string)
+ | _ => returnm None
end)
: M (option string)
end)
: M (option string).
-Definition sep_matches_prefix (arg_ : string)
-: M (option ((unit * {n : Z & ArithFact (n >= 0)}))) :=
-
+Definition sep_matches_prefix (arg_ : string) : M (option ((unit * {n : Z & ArithFact (n >=? 0)}))) :=
let _s23_ := arg_ in
(_s16_ _s23_) >>= fun w__0 : option string =>
- (if ((match w__0 with | Some (s_) => true | _ => false end)) then
+ (if match w__0 with | Some s_ => true | _ => false end then
(_s16_ _s23_) >>= fun w__1 : option string =>
(match w__1 with
- | Some (s_) =>
- returnm ((Some
- ((tt, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((unit * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((unit * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ (tt, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((unit * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((unit * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((unit * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((unit * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((unit * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((unit * {n : Z & ArithFact (n >=? 0)}))).
-Definition bool_bits_forwards (arg_ : bool)
-: mword 1 :=
-
- match arg_ with
- | true => (vec_of_bits [B1] : mword 1)
- | false => (vec_of_bits [B0] : mword 1)
- end.
+Definition bool_bits_forwards (arg_ : bool) : mword 1 :=
+ match arg_ with | true => 'b"1" : mword 1 | false => 'b"0" : mword 1 end.
-Definition bool_bits_backwards (arg_ : mword 1)
-: M (bool) :=
-
+Definition bool_bits_backwards (arg_ : mword 1) : M (bool) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ (if eq_vec b__0 ('b"1" : mword 1) then returnm true
+ else if eq_vec b__0 ('b"0" : mword 1) then returnm false
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bool).
-Definition bool_bits_forwards_matches (arg_ : bool)
-: bool :=
-
+Definition bool_bits_forwards_matches (arg_ : bool) : bool :=
match arg_ with | true => true | false => true end.
-Definition bool_bits_backwards_matches (arg_ : mword 1)
-: bool :=
-
+Definition bool_bits_backwards_matches (arg_ : mword 1) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
+ if eq_vec b__0 ('b"1" : mword 1) then true
+ else if eq_vec b__0 ('b"0" : mword 1) then true
else false.
-Definition bool_not_bits_forwards (arg_ : bool)
-: mword 1 :=
-
- match arg_ with
- | true => (vec_of_bits [B0] : mword 1)
- | false => (vec_of_bits [B1] : mword 1)
- end.
+Definition bool_not_bits_forwards (arg_ : bool) : mword 1 :=
+ match arg_ with | true => 'b"0" : mword 1 | false => 'b"1" : mword 1 end.
-Definition bool_not_bits_backwards (arg_ : mword 1)
-: M (bool) :=
-
+Definition bool_not_bits_backwards (arg_ : mword 1) : M (bool) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ (if eq_vec b__0 ('b"0" : mword 1) then returnm true
+ else if eq_vec b__0 ('b"1" : mword 1) then returnm false
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bool).
-Definition bool_not_bits_forwards_matches (arg_ : bool)
-: bool :=
-
+Definition bool_not_bits_forwards_matches (arg_ : bool) : bool :=
match arg_ with | true => true | false => true end.
-Definition bool_not_bits_backwards_matches (arg_ : mword 1)
-: bool :=
-
+Definition bool_not_bits_backwards_matches (arg_ : mword 1) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
+ if eq_vec b__0 ('b"0" : mword 1) then true
+ else if eq_vec b__0 ('b"1" : mword 1) then true
else false.
-Definition size_bits_forwards (arg_ : word_width)
-: mword 2 :=
-
+Definition size_bits_forwards (arg_ : word_width) : mword 2 :=
match arg_ with
- | BYTE => (vec_of_bits [B0;B0] : mword 2)
- | HALF => (vec_of_bits [B0;B1] : mword 2)
- | WORD => (vec_of_bits [B1;B0] : mword 2)
- | DOUBLE => (vec_of_bits [B1;B1] : mword 2)
+ | BYTE => 'b"00" : mword 2
+ | HALF => 'b"01" : mword 2
+ | WORD => 'b"10" : mword 2
+ | DOUBLE => 'b"11" : mword 2
end.
-Definition size_bits_backwards (arg_ : mword 2)
-: M (word_width) :=
-
+Definition size_bits_backwards (arg_ : mword 2) : M (word_width) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0] : mword 2))) then returnm (BYTE : word_width)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then returnm (HALF : word_width)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then returnm (WORD : word_width)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then returnm (DOUBLE : word_width)
+ (if eq_vec b__0 ('b"00" : mword 2) then returnm BYTE
+ else if eq_vec b__0 ('b"01" : mword 2) then returnm HALF
+ else if eq_vec b__0 ('b"10" : mword 2) then returnm WORD
+ else if eq_vec b__0 ('b"11" : mword 2) then returnm DOUBLE
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (word_width).
-Definition size_bits_forwards_matches (arg_ : word_width)
-: bool :=
-
+Definition size_bits_forwards_matches (arg_ : word_width) : bool :=
match arg_ with | BYTE => true | HALF => true | WORD => true | DOUBLE => true end.
-Definition size_bits_backwards_matches (arg_ : mword 2)
-: bool :=
-
+Definition size_bits_backwards_matches (arg_ : mword 2) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0] : mword 2))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then true
+ if eq_vec b__0 ('b"00" : mword 2) then true
+ else if eq_vec b__0 ('b"01" : mword 2) then true
+ else if eq_vec b__0 ('b"10" : mword 2) then true
+ else if eq_vec b__0 ('b"11" : mword 2) then true
else false.
-Definition size_mnemonic_forwards (arg_ : word_width)
-: string :=
-
+Definition size_mnemonic_forwards (arg_ : word_width) : string :=
match arg_ with | BYTE => "b" | HALF => "h" | WORD => "w" | DOUBLE => "d" end.
-Definition size_mnemonic_backwards (arg_ : string)
-: M (word_width) :=
-
+Definition size_mnemonic_backwards (arg_ : string) : M (word_width) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "b")) then returnm (BYTE : word_width)
- else if ((generic_eq p0_ "h")) then returnm (HALF : word_width)
- else if ((generic_eq p0_ "w")) then returnm (WORD : word_width)
- else if ((generic_eq p0_ "d")) then returnm (DOUBLE : word_width)
+ (if generic_eq p0_ "b" then returnm BYTE
+ else if generic_eq p0_ "h" then returnm HALF
+ else if generic_eq p0_ "w" then returnm WORD
+ else if generic_eq p0_ "d" then returnm DOUBLE
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (word_width).
-Definition size_mnemonic_forwards_matches (arg_ : word_width)
-: bool :=
-
+Definition size_mnemonic_forwards_matches (arg_ : word_width) : bool :=
match arg_ with | BYTE => true | HALF => true | WORD => true | DOUBLE => true end.
-Definition size_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition size_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "b")) then true
- else if ((generic_eq p0_ "h")) then true
- else if ((generic_eq p0_ "w")) then true
- else if ((generic_eq p0_ "d")) then true
+ if generic_eq p0_ "b" then true
+ else if generic_eq p0_ "h" then true
+ else if generic_eq p0_ "w" then true
+ else if generic_eq p0_ "d" then true
else false.
-Definition _s36_ (_s37_ : string)
-: option string :=
-
+Definition _s36_ (_s37_ : string) : option string :=
let _s38_ := _s37_ in
- if ((string_startswith _s38_ "d")) then
- match (string_drop _s38_ (projT1 (string_length "d"))) with | s_ => Some (s_) end
+ if string_startswith _s38_ "d" then
+ match (string_drop _s38_ (projT1 (string_length "d"))) with | s_ => Some s_ end
else None.
-Definition _s32_ (_s33_ : string)
-: option string :=
-
+Definition _s32_ (_s33_ : string) : option string :=
let _s34_ := _s33_ in
- if ((string_startswith _s34_ "w")) then
- match (string_drop _s34_ (projT1 (string_length "w"))) with | s_ => Some (s_) end
+ if string_startswith _s34_ "w" then
+ match (string_drop _s34_ (projT1 (string_length "w"))) with | s_ => Some s_ end
else None.
-Definition _s28_ (_s29_ : string)
-: option string :=
-
+Definition _s28_ (_s29_ : string) : option string :=
let _s30_ := _s29_ in
- if ((string_startswith _s30_ "h")) then
- match (string_drop _s30_ (projT1 (string_length "h"))) with | s_ => Some (s_) end
+ if string_startswith _s30_ "h" then
+ match (string_drop _s30_ (projT1 (string_length "h"))) with | s_ => Some s_ end
else None.
-Definition _s24_ (_s25_ : string)
-: option string :=
-
+Definition _s24_ (_s25_ : string) : option string :=
let _s26_ := _s25_ in
- if ((string_startswith _s26_ "b")) then
- match (string_drop _s26_ (projT1 (string_length "b"))) with | s_ => Some (s_) end
+ if string_startswith _s26_ "b" then
+ match (string_drop _s26_ (projT1 (string_length "b"))) with | s_ => Some s_ end
else None.
-Definition size_mnemonic_matches_prefix (arg_ : string)
-: M (option ((word_width * {n : Z & ArithFact (n >= 0)}))) :=
-
+Definition size_mnemonic_matches_prefix (arg_ : string)
+: M (option ((word_width * {n : Z & ArithFact (n >=? 0)}))) :=
let _s27_ := arg_ in
- (if ((match (_s24_ _s27_) with | Some (s_) => true | _ => false end)) then
+ (if match (_s24_ _s27_) with | Some s_ => true | _ => false end then
(match (_s24_ _s27_) with
- | Some (s_) =>
- returnm ((Some
- ((BYTE, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((word_width * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((word_width * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ (BYTE, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((word_width * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((word_width * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s28_ _s27_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((word_width * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s28_ _s27_) with | Some s_ => true | _ => false end then
(match (_s28_ _s27_) with
- | Some (s_) =>
- returnm ((Some
- ((HALF, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((word_width * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((word_width * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ (HALF, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((word_width * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((word_width * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s32_ _s27_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((word_width * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s32_ _s27_) with | Some s_ => true | _ => false end then
(match (_s32_ _s27_) with
- | Some (s_) =>
- returnm ((Some
- ((WORD, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((word_width * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((word_width * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ (WORD, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((word_width * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((word_width * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s36_ _s27_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((word_width * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s36_ _s27_) with | Some s_ => true | _ => false end then
(match (_s36_ _s27_) with
- | Some (s_) =>
- returnm ((Some
- ((DOUBLE, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((word_width * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((word_width * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ (DOUBLE, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((word_width * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((word_width * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((word_width * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((word_width * {n : Z & ArithFact (n >= 0)}))).
-
-Definition word_width_bytes (width : word_width)
-: {s : Z & ArithFact (s = 1 \/ s = 2 \/ s = 4 \/ s = 8)} :=
-
- build_ex(match width with | BYTE => 1 | HALF => 2 | WORD => 4 | DOUBLE => 8 end).
+ : M (option ((word_width * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((word_width * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition word_width_bytes (width : word_width)
+: {s : Z & ArithFact ((s =? 1) || ((s =? 2) || ((s =? 4) || (s =? 8))))} :=
+ build_ex (match width with | BYTE => 1 | HALF => 2 | WORD => 4 | DOUBLE => 8 end).
+
+Definition Data : ext_access_type := tt.
+Hint Unfold Data : sail.
+Definition default_write_acc : ext_access_type := Data.
+Hint Unfold default_write_acc : sail.
+Definition accessType_to_str (a : AccessType unit) : string :=
+ match a with
+ | Read Data => "R"
+ | Write Data => "W"
+ | ReadWrite Data => "RW"
+ | Execute tt => "X"
+ end.
-Definition zero_reg : regtype := EXTZ 64 (vec_of_bits [B0;B0;B0;B0] : mword 4).
+Definition zero_reg : regtype := EXTZ 64 (Ox"0" : mword 4).
Hint Unfold zero_reg : sail.
-Definition RegStr (r : mword 64) : string := string_of_bits r.
-
-Definition regval_from_reg (r : mword 64) : mword 64 := r.
+Definition RegStr (r : mword 64) : string := string_of_bits r.
+
+Definition regval_from_reg (r : mword 64) : mword 64 := r.
+
+Definition regval_into_reg (v : mword 64) : mword 64 := v.
+
+Definition zero_freg : fregtype := EXTZ 64 (Ox"0" : mword 4).
+Hint Unfold zero_freg : sail.
+Definition FRegStr (r : mword 64) : string := string_of_bits r.
+
+Definition fregval_from_freg (r : mword 64) : mword 64 := r.
+
+Definition fregval_into_freg (v : mword 64) : mword 64 := v.
+
+Definition rounding_mode_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 5))}
+: rounding_mode :=
+ let l__197 := arg_ in
+ if sumbool_of_bool (Z.eqb l__197 0) then RM_RNE
+ else if sumbool_of_bool (Z.eqb l__197 1) then RM_RTZ
+ else if sumbool_of_bool (Z.eqb l__197 2) then RM_RDN
+ else if sumbool_of_bool (Z.eqb l__197 3) then RM_RUP
+ else if sumbool_of_bool (Z.eqb l__197 4) then RM_RMM
+ else RM_DYN.
+
+Definition num_of_rounding_mode (arg_ : rounding_mode)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 5))} :=
+ build_ex (
+ match arg_ with
+ | RM_RNE => 0
+ | RM_RTZ => 1
+ | RM_RDN => 2
+ | RM_RUP => 3
+ | RM_RMM => 4
+ | RM_DYN => 5
+ end
+ ).
+
+Definition f_madd_op_S_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))} : f_madd_op_S :=
+ let l__194 := arg_ in
+ if sumbool_of_bool (Z.eqb l__194 0) then FMADD_S
+ else if sumbool_of_bool (Z.eqb l__194 1) then FMSUB_S
+ else if sumbool_of_bool (Z.eqb l__194 2) then FNMSUB_S
+ else FNMADD_S.
+
+Definition num_of_f_madd_op_S (arg_ : f_madd_op_S) : {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (match arg_ with | FMADD_S => 0 | FMSUB_S => 1 | FNMSUB_S => 2 | FNMADD_S => 3 end).
+
+Definition f_bin_rm_op_S_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))}
+: f_bin_rm_op_S :=
+ let l__191 := arg_ in
+ if sumbool_of_bool (Z.eqb l__191 0) then FADD_S
+ else if sumbool_of_bool (Z.eqb l__191 1) then FSUB_S
+ else if sumbool_of_bool (Z.eqb l__191 2) then FMUL_S
+ else FDIV_S.
+
+Definition num_of_f_bin_rm_op_S (arg_ : f_bin_rm_op_S)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (match arg_ with | FADD_S => 0 | FSUB_S => 1 | FMUL_S => 2 | FDIV_S => 3 end).
+
+Definition f_un_rm_op_S_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 8))} : f_un_rm_op_S :=
+ let l__183 := arg_ in
+ if sumbool_of_bool (Z.eqb l__183 0) then FSQRT_S
+ else if sumbool_of_bool (Z.eqb l__183 1) then FCVT_W_S
+ else if sumbool_of_bool (Z.eqb l__183 2) then FCVT_WU_S
+ else if sumbool_of_bool (Z.eqb l__183 3) then FCVT_S_W
+ else if sumbool_of_bool (Z.eqb l__183 4) then FCVT_S_WU
+ else if sumbool_of_bool (Z.eqb l__183 5) then FCVT_L_S
+ else if sumbool_of_bool (Z.eqb l__183 6) then FCVT_LU_S
+ else if sumbool_of_bool (Z.eqb l__183 7) then FCVT_S_L
+ else FCVT_S_LU.
+
+Definition num_of_f_un_rm_op_S (arg_ : f_un_rm_op_S) : {e : Z & ArithFact ((0 <=? e) && (e <=? 8))} :=
+ build_ex (
+ match arg_ with
+ | FSQRT_S => 0
+ | FCVT_W_S => 1
+ | FCVT_WU_S => 2
+ | FCVT_S_W => 3
+ | FCVT_S_WU => 4
+ | FCVT_L_S => 5
+ | FCVT_LU_S => 6
+ | FCVT_S_L => 7
+ | FCVT_S_LU => 8
+ end
+ ).
+
+Definition f_un_op_S_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : f_un_op_S :=
+ let l__181 := arg_ in
+ if sumbool_of_bool (Z.eqb l__181 0) then FCLASS_S
+ else if sumbool_of_bool (Z.eqb l__181 1) then FMV_X_W
+ else FMV_W_X.
+
+Definition num_of_f_un_op_S (arg_ : f_un_op_S) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | FCLASS_S => 0 | FMV_X_W => 1 | FMV_W_X => 2 end).
+
+Definition f_bin_op_S_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 7))} : f_bin_op_S :=
+ let l__174 := arg_ in
+ if sumbool_of_bool (Z.eqb l__174 0) then FSGNJ_S
+ else if sumbool_of_bool (Z.eqb l__174 1) then FSGNJN_S
+ else if sumbool_of_bool (Z.eqb l__174 2) then FSGNJX_S
+ else if sumbool_of_bool (Z.eqb l__174 3) then FMIN_S
+ else if sumbool_of_bool (Z.eqb l__174 4) then FMAX_S
+ else if sumbool_of_bool (Z.eqb l__174 5) then FEQ_S
+ else if sumbool_of_bool (Z.eqb l__174 6) then FLT_S
+ else FLE_S.
+
+Definition num_of_f_bin_op_S (arg_ : f_bin_op_S) : {e : Z & ArithFact ((0 <=? e) && (e <=? 7))} :=
+ build_ex (
+ match arg_ with
+ | FSGNJ_S => 0
+ | FSGNJN_S => 1
+ | FSGNJX_S => 2
+ | FMIN_S => 3
+ | FMAX_S => 4
+ | FEQ_S => 5
+ | FLT_S => 6
+ | FLE_S => 7
+ end
+ ).
+
+Definition f_madd_op_D_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))} : f_madd_op_D :=
+ let l__171 := arg_ in
+ if sumbool_of_bool (Z.eqb l__171 0) then FMADD_D
+ else if sumbool_of_bool (Z.eqb l__171 1) then FMSUB_D
+ else if sumbool_of_bool (Z.eqb l__171 2) then FNMSUB_D
+ else FNMADD_D.
+
+Definition num_of_f_madd_op_D (arg_ : f_madd_op_D) : {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (match arg_ with | FMADD_D => 0 | FMSUB_D => 1 | FNMSUB_D => 2 | FNMADD_D => 3 end).
+
+Definition f_bin_rm_op_D_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))}
+: f_bin_rm_op_D :=
+ let l__168 := arg_ in
+ if sumbool_of_bool (Z.eqb l__168 0) then FADD_D
+ else if sumbool_of_bool (Z.eqb l__168 1) then FSUB_D
+ else if sumbool_of_bool (Z.eqb l__168 2) then FMUL_D
+ else FDIV_D.
+
+Definition num_of_f_bin_rm_op_D (arg_ : f_bin_rm_op_D)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (match arg_ with | FADD_D => 0 | FSUB_D => 1 | FMUL_D => 2 | FDIV_D => 3 end).
+
+Definition f_un_rm_op_D_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 10))}
+: f_un_rm_op_D :=
+ let l__158 := arg_ in
+ if sumbool_of_bool (Z.eqb l__158 0) then FSQRT_D
+ else if sumbool_of_bool (Z.eqb l__158 1) then FCVT_W_D
+ else if sumbool_of_bool (Z.eqb l__158 2) then FCVT_WU_D
+ else if sumbool_of_bool (Z.eqb l__158 3) then FCVT_D_W
+ else if sumbool_of_bool (Z.eqb l__158 4) then FCVT_D_WU
+ else if sumbool_of_bool (Z.eqb l__158 5) then FCVT_S_D
+ else if sumbool_of_bool (Z.eqb l__158 6) then FCVT_D_S
+ else if sumbool_of_bool (Z.eqb l__158 7) then FCVT_L_D
+ else if sumbool_of_bool (Z.eqb l__158 8) then FCVT_LU_D
+ else if sumbool_of_bool (Z.eqb l__158 9) then FCVT_D_L
+ else FCVT_D_LU.
+
+Definition num_of_f_un_rm_op_D (arg_ : f_un_rm_op_D) : {e : Z & ArithFact ((0 <=? e) && (e <=? 10))} :=
+ build_ex (
+ match arg_ with
+ | FSQRT_D => 0
+ | FCVT_W_D => 1
+ | FCVT_WU_D => 2
+ | FCVT_D_W => 3
+ | FCVT_D_WU => 4
+ | FCVT_S_D => 5
+ | FCVT_D_S => 6
+ | FCVT_L_D => 7
+ | FCVT_LU_D => 8
+ | FCVT_D_L => 9
+ | FCVT_D_LU => 10
+ end
+ ).
+
+Definition f_bin_op_D_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 7))} : f_bin_op_D :=
+ let l__151 := arg_ in
+ if sumbool_of_bool (Z.eqb l__151 0) then FSGNJ_D
+ else if sumbool_of_bool (Z.eqb l__151 1) then FSGNJN_D
+ else if sumbool_of_bool (Z.eqb l__151 2) then FSGNJX_D
+ else if sumbool_of_bool (Z.eqb l__151 3) then FMIN_D
+ else if sumbool_of_bool (Z.eqb l__151 4) then FMAX_D
+ else if sumbool_of_bool (Z.eqb l__151 5) then FEQ_D
+ else if sumbool_of_bool (Z.eqb l__151 6) then FLT_D
+ else FLE_D.
+
+Definition num_of_f_bin_op_D (arg_ : f_bin_op_D) : {e : Z & ArithFact ((0 <=? e) && (e <=? 7))} :=
+ build_ex (
+ match arg_ with
+ | FSGNJ_D => 0
+ | FSGNJN_D => 1
+ | FSGNJX_D => 2
+ | FMIN_D => 3
+ | FMAX_D => 4
+ | FEQ_D => 5
+ | FLT_D => 6
+ | FLE_D => 7
+ end
+ ).
+
+Definition f_un_op_D_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : f_un_op_D :=
+ let l__149 := arg_ in
+ if sumbool_of_bool (Z.eqb l__149 0) then FCLASS_D
+ else if sumbool_of_bool (Z.eqb l__149 1) then FMV_X_D
+ else FMV_D_X.
-Definition regval_into_reg (v : mword 64) : mword 64 := v.
+Definition num_of_f_un_op_D (arg_ : f_un_op_D) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | FCLASS_D => 0 | FMV_X_D => 1 | FMV_D_X => 2 end).
-Definition rX (r : Z) `{ArithFact (0 <= r /\ r < 32)}
-: M (mword 64) :=
-
- let l__55 := r in
- (if sumbool_of_bool ((Z.eqb l__55 0)) then returnm (zero_reg : mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 1)) then
+Definition rX (r : Z) `{ArithFact ((0 <=? r) && (r <? 32))} : M (mword 64) :=
+ let l__117 := r in
+ (if sumbool_of_bool (Z.eqb l__117 0) then returnm zero_reg
+ else if sumbool_of_bool (Z.eqb l__117 1) then
((read_reg x1_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 2)) then
+ else if sumbool_of_bool (Z.eqb l__117 2) then
((read_reg x2_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 3)) then
+ else if sumbool_of_bool (Z.eqb l__117 3) then
((read_reg x3_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 4)) then
+ else if sumbool_of_bool (Z.eqb l__117 4) then
((read_reg x4_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 5)) then
+ else if sumbool_of_bool (Z.eqb l__117 5) then
((read_reg x5_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 6)) then
+ else if sumbool_of_bool (Z.eqb l__117 6) then
((read_reg x6_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 7)) then
+ else if sumbool_of_bool (Z.eqb l__117 7) then
((read_reg x7_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 8)) then
+ else if sumbool_of_bool (Z.eqb l__117 8) then
((read_reg x8_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 9)) then
+ else if sumbool_of_bool (Z.eqb l__117 9) then
((read_reg x9_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 10)) then
+ else if sumbool_of_bool (Z.eqb l__117 10) then
((read_reg x10_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 11)) then
+ else if sumbool_of_bool (Z.eqb l__117 11) then
((read_reg x11_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 12)) then
+ else if sumbool_of_bool (Z.eqb l__117 12) then
((read_reg x12_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 13)) then
+ else if sumbool_of_bool (Z.eqb l__117 13) then
((read_reg x13_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 14)) then
+ else if sumbool_of_bool (Z.eqb l__117 14) then
((read_reg x14_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 15)) then
+ else if sumbool_of_bool (Z.eqb l__117 15) then
((read_reg x15_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 16)) then
+ else if sumbool_of_bool (Z.eqb l__117 16) then
((read_reg x16_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 17)) then
+ else if sumbool_of_bool (Z.eqb l__117 17) then
((read_reg x17_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 18)) then
+ else if sumbool_of_bool (Z.eqb l__117 18) then
((read_reg x18_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 19)) then
+ else if sumbool_of_bool (Z.eqb l__117 19) then
((read_reg x19_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 20)) then
+ else if sumbool_of_bool (Z.eqb l__117 20) then
((read_reg x20_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 21)) then
+ else if sumbool_of_bool (Z.eqb l__117 21) then
((read_reg x21_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 22)) then
+ else if sumbool_of_bool (Z.eqb l__117 22) then
((read_reg x22_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 23)) then
+ else if sumbool_of_bool (Z.eqb l__117 23) then
((read_reg x23_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 24)) then
+ else if sumbool_of_bool (Z.eqb l__117 24) then
((read_reg x24_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 25)) then
+ else if sumbool_of_bool (Z.eqb l__117 25) then
((read_reg x25_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 26)) then
+ else if sumbool_of_bool (Z.eqb l__117 26) then
((read_reg x26_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 27)) then
+ else if sumbool_of_bool (Z.eqb l__117 27) then
((read_reg x27_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 28)) then
+ else if sumbool_of_bool (Z.eqb l__117 28) then
((read_reg x28_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 29)) then
+ else if sumbool_of_bool (Z.eqb l__117 29) then
((read_reg x29_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 30)) then
+ else if sumbool_of_bool (Z.eqb l__117 30) then
((read_reg x30_ref) : M (mword 64))
: M (mword 64)
- else if sumbool_of_bool ((Z.eqb l__55 31)) then
+ else if sumbool_of_bool (Z.eqb l__117 31) then
((read_reg x31_ref) : M (mword 64))
: M (mword 64)
else assert_exp' false "invalid register number" >>= fun _ => exit tt) >>= fun v : regtype =>
- returnm ((regval_from_reg v)
- : mword 64).
+ returnm (regval_from_reg v).
-Definition rvfi_wX (r : Z) (v : mword 64) `{ArithFact (0 <= r /\ r < 32)} : unit := tt.
+Definition rvfi_wX (r : Z) (v : mword 64) `{ArithFact ((0 <=? r) && (r <? 32))} : unit := tt.
-Definition wX (r : Z) (in_v : mword 64) `{ArithFact (0 <= r /\ r < 32)}
-: M (unit) :=
-
+Definition wX (r : Z) (in_v : mword 64) `{ArithFact ((0 <=? r) && (r <? 32))} : M (unit) :=
let v := regval_into_reg in_v in
- let l__23 := r in
- (if sumbool_of_bool ((Z.eqb l__23 0)) then returnm (tt : unit)
- else if sumbool_of_bool ((Z.eqb l__23 1)) then write_reg x1_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 2)) then write_reg x2_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 3)) then write_reg x3_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 4)) then write_reg x4_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 5)) then write_reg x5_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 6)) then write_reg x6_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 7)) then write_reg x7_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 8)) then write_reg x8_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 9)) then write_reg x9_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 10)) then write_reg x10_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 11)) then write_reg x11_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 12)) then write_reg x12_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 13)) then write_reg x13_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 14)) then write_reg x14_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 15)) then write_reg x15_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 16)) then write_reg x16_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 17)) then write_reg x17_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 18)) then write_reg x18_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 19)) then write_reg x19_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 20)) then write_reg x20_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 21)) then write_reg x21_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 22)) then write_reg x22_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 23)) then write_reg x23_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 24)) then write_reg x24_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 25)) then write_reg x25_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 26)) then write_reg x26_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 27)) then write_reg x27_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 28)) then write_reg x28_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 29)) then write_reg x29_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 30)) then write_reg x30_ref v : M (unit)
- else if sumbool_of_bool ((Z.eqb l__23 31)) then write_reg x31_ref v : M (unit)
+ let l__85 := r in
+ (if sumbool_of_bool (Z.eqb l__85 0) then returnm tt
+ else if sumbool_of_bool (Z.eqb l__85 1) then write_reg x1_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 2) then write_reg x2_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 3) then write_reg x3_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 4) then write_reg x4_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 5) then write_reg x5_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 6) then write_reg x6_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 7) then write_reg x7_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 8) then write_reg x8_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 9) then write_reg x9_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 10) then write_reg x10_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 11) then write_reg x11_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 12) then write_reg x12_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 13) then write_reg x13_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 14) then write_reg x14_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 15) then write_reg x15_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 16) then write_reg x16_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 17) then write_reg x17_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 18) then write_reg x18_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 19) then write_reg x19_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 20) then write_reg x20_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 21) then write_reg x21_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 22) then write_reg x22_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 23) then write_reg x23_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 24) then write_reg x24_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 25) then write_reg x25_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 26) then write_reg x26_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 27) then write_reg x27_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 28) then write_reg x28_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 29) then write_reg x29_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 30) then write_reg x30_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__85 31) then write_reg x31_ref v : M (unit)
else assert_exp' false "invalid register number" >>= fun _ => exit tt) >>
- returnm ((if sumbool_of_bool ((projT1 (neq_int r 0))) then
- let '_ := (rvfi_wX r in_v) : unit in
- if ((get_config_print_reg tt)) then
- print_endline
- (String.append "x"
- (String.append (string_of_int r) (String.append " <- " (RegStr v))))
- else tt
- else tt)
- : unit).
+ returnm (if sumbool_of_bool (projT1 (neq_int r 0)) then
+ let '_ := (rvfi_wX r in_v) : unit in
+ if get_config_print_reg tt then
+ print_endline
+ (String.append "x"
+ (String.append (string_of_int r) (String.append " <- " (RegStr v))))
+ else tt
+ else tt).
+
+Definition rX_bits (i : mword 5) : M (mword 64) := (rX (projT1 (uint i))) : M (mword 64).
+
+Definition wX_bits (i : mword 5) (data : mword 64) : M (unit) :=
+ (wX (projT1 (uint i)) data) : M (unit).
-Definition reg_name_abi (r : mword 5)
-: M (string) :=
-
+Definition reg_name_abi (r : mword 5) : M (string) :=
let b__0 := r in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))) then returnm ("zero" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))) then returnm ("ra" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))) then returnm ("sp" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1] : mword 5))) then returnm ("gp" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))) then returnm ("tp" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B1] : mword 5))) then returnm ("t0" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0] : mword 5))) then returnm ("t1" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1] : mword 5))) then returnm ("t2" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))) then returnm ("fp" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B0;B1] : mword 5))) then returnm ("s1" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B1;B0] : mword 5))) then returnm ("a0" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B1;B1] : mword 5))) then returnm ("a1" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))) then returnm ("a2" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B1] : mword 5))) then returnm ("a3" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B0] : mword 5))) then returnm ("a4" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1] : mword 5))) then returnm ("a5" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))) then returnm ("a6" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0;B1] : mword 5))) then returnm ("a7" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B1;B0] : mword 5))) then returnm ("s2" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B1;B1] : mword 5))) then returnm ("s3" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))) then returnm ("s4" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B0;B1] : mword 5))) then returnm ("s5" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0] : mword 5))) then returnm ("s6" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1] : mword 5))) then returnm ("s7" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))) then returnm ("s8" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1] : mword 5))) then returnm ("s9" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B1;B0] : mword 5))) then returnm ("s10" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B1;B1] : mword 5))) then returnm ("s11" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))) then returnm ("t3" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B0;B1] : mword 5))) then returnm ("t4" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0] : mword 5))) then returnm ("t5" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B1] : mword 5))) then returnm ("t6" : string)
+ (if eq_vec b__0 ('b"00000" : mword 5) then returnm "zero"
+ else if eq_vec b__0 ('b"00001" : mword 5) then returnm "ra"
+ else if eq_vec b__0 ('b"00010" : mword 5) then returnm "sp"
+ else if eq_vec b__0 ('b"00011" : mword 5) then returnm "gp"
+ else if eq_vec b__0 ('b"00100" : mword 5) then returnm "tp"
+ else if eq_vec b__0 ('b"00101" : mword 5) then returnm "t0"
+ else if eq_vec b__0 ('b"00110" : mword 5) then returnm "t1"
+ else if eq_vec b__0 ('b"00111" : mword 5) then returnm "t2"
+ else if eq_vec b__0 ('b"01000" : mword 5) then returnm "fp"
+ else if eq_vec b__0 ('b"01001" : mword 5) then returnm "s1"
+ else if eq_vec b__0 ('b"01010" : mword 5) then returnm "a0"
+ else if eq_vec b__0 ('b"01011" : mword 5) then returnm "a1"
+ else if eq_vec b__0 ('b"01100" : mword 5) then returnm "a2"
+ else if eq_vec b__0 ('b"01101" : mword 5) then returnm "a3"
+ else if eq_vec b__0 ('b"01110" : mword 5) then returnm "a4"
+ else if eq_vec b__0 ('b"01111" : mword 5) then returnm "a5"
+ else if eq_vec b__0 ('b"10000" : mword 5) then returnm "a6"
+ else if eq_vec b__0 ('b"10001" : mword 5) then returnm "a7"
+ else if eq_vec b__0 ('b"10010" : mword 5) then returnm "s2"
+ else if eq_vec b__0 ('b"10011" : mword 5) then returnm "s3"
+ else if eq_vec b__0 ('b"10100" : mword 5) then returnm "s4"
+ else if eq_vec b__0 ('b"10101" : mword 5) then returnm "s5"
+ else if eq_vec b__0 ('b"10110" : mword 5) then returnm "s6"
+ else if eq_vec b__0 ('b"10111" : mword 5) then returnm "s7"
+ else if eq_vec b__0 ('b"11000" : mword 5) then returnm "s8"
+ else if eq_vec b__0 ('b"11001" : mword 5) then returnm "s9"
+ else if eq_vec b__0 ('b"11010" : mword 5) then returnm "s10"
+ else if eq_vec b__0 ('b"11011" : mword 5) then returnm "s11"
+ else if eq_vec b__0 ('b"11100" : mword 5) then returnm "t3"
+ else if eq_vec b__0 ('b"11101" : mword 5) then returnm "t4"
+ else if eq_vec b__0 ('b"11110" : mword 5) then returnm "t5"
+ else if eq_vec b__0 ('b"11111" : mword 5) then returnm "t6"
else
- assert_exp' false "Pattern match failure at model/riscv_regs.sail 149:2 - 182:3" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/riscv_regs.sail 155:2 - 188:3" >>= fun _ =>
exit tt)
: M (string).
-Definition reg_name_forwards (arg_ : mword 5)
-: M (string) :=
-
+Definition reg_name_forwards (arg_ : mword 5) : M (string) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))) then returnm ("zero" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))) then returnm ("ra" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))) then returnm ("sp" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1] : mword 5))) then returnm ("gp" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))) then returnm ("tp" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B1] : mword 5))) then returnm ("t0" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0] : mword 5))) then returnm ("t1" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1] : mword 5))) then returnm ("t2" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))) then returnm ("fp" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B0;B1] : mword 5))) then returnm ("s1" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B1;B0] : mword 5))) then returnm ("a0" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B1;B1] : mword 5))) then returnm ("a1" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))) then returnm ("a2" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B1] : mword 5))) then returnm ("a3" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B0] : mword 5))) then returnm ("a4" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1] : mword 5))) then returnm ("a5" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))) then returnm ("a6" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0;B1] : mword 5))) then returnm ("a7" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B1;B0] : mword 5))) then returnm ("s2" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B1;B1] : mword 5))) then returnm ("s3" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))) then returnm ("s4" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B0;B1] : mword 5))) then returnm ("s5" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0] : mword 5))) then returnm ("s6" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1] : mword 5))) then returnm ("s7" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))) then returnm ("s8" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1] : mword 5))) then returnm ("s9" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B1;B0] : mword 5))) then returnm ("s10" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B1;B1] : mword 5))) then returnm ("s11" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))) then returnm ("t3" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B0;B1] : mword 5))) then returnm ("t4" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0] : mword 5))) then returnm ("t5" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B1] : mword 5))) then returnm ("t6" : string)
+ (if eq_vec b__0 ('b"00000" : mword 5) then returnm "zero"
+ else if eq_vec b__0 ('b"00001" : mword 5) then returnm "ra"
+ else if eq_vec b__0 ('b"00010" : mword 5) then returnm "sp"
+ else if eq_vec b__0 ('b"00011" : mword 5) then returnm "gp"
+ else if eq_vec b__0 ('b"00100" : mword 5) then returnm "tp"
+ else if eq_vec b__0 ('b"00101" : mword 5) then returnm "t0"
+ else if eq_vec b__0 ('b"00110" : mword 5) then returnm "t1"
+ else if eq_vec b__0 ('b"00111" : mword 5) then returnm "t2"
+ else if eq_vec b__0 ('b"01000" : mword 5) then returnm "fp"
+ else if eq_vec b__0 ('b"01001" : mword 5) then returnm "s1"
+ else if eq_vec b__0 ('b"01010" : mword 5) then returnm "a0"
+ else if eq_vec b__0 ('b"01011" : mword 5) then returnm "a1"
+ else if eq_vec b__0 ('b"01100" : mword 5) then returnm "a2"
+ else if eq_vec b__0 ('b"01101" : mword 5) then returnm "a3"
+ else if eq_vec b__0 ('b"01110" : mword 5) then returnm "a4"
+ else if eq_vec b__0 ('b"01111" : mword 5) then returnm "a5"
+ else if eq_vec b__0 ('b"10000" : mword 5) then returnm "a6"
+ else if eq_vec b__0 ('b"10001" : mword 5) then returnm "a7"
+ else if eq_vec b__0 ('b"10010" : mword 5) then returnm "s2"
+ else if eq_vec b__0 ('b"10011" : mword 5) then returnm "s3"
+ else if eq_vec b__0 ('b"10100" : mword 5) then returnm "s4"
+ else if eq_vec b__0 ('b"10101" : mword 5) then returnm "s5"
+ else if eq_vec b__0 ('b"10110" : mword 5) then returnm "s6"
+ else if eq_vec b__0 ('b"10111" : mword 5) then returnm "s7"
+ else if eq_vec b__0 ('b"11000" : mword 5) then returnm "s8"
+ else if eq_vec b__0 ('b"11001" : mword 5) then returnm "s9"
+ else if eq_vec b__0 ('b"11010" : mword 5) then returnm "s10"
+ else if eq_vec b__0 ('b"11011" : mword 5) then returnm "s11"
+ else if eq_vec b__0 ('b"11100" : mword 5) then returnm "t3"
+ else if eq_vec b__0 ('b"11101" : mword 5) then returnm "t4"
+ else if eq_vec b__0 ('b"11110" : mword 5) then returnm "t5"
+ else if eq_vec b__0 ('b"11111" : mword 5) then returnm "t6"
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string).
-Definition reg_name_backwards (arg_ : string)
-: M (mword 5) :=
-
+Definition reg_name_backwards (arg_ : string) : M (mword 5) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "zero")) then returnm ((vec_of_bits [B0;B0;B0;B0;B0] : mword 5) : mword 5)
- else if ((generic_eq p0_ "ra")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "sp")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "gp")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "tp")) then
- returnm ((vec_of_bits [B0;B0;B1;B0;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "t0")) then
- returnm ((vec_of_bits [B0;B0;B1;B0;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "t1")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "t2")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "fp")) then
- returnm ((vec_of_bits [B0;B1;B0;B0;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s1")) then
- returnm ((vec_of_bits [B0;B1;B0;B0;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "a0")) then
- returnm ((vec_of_bits [B0;B1;B0;B1;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "a1")) then
- returnm ((vec_of_bits [B0;B1;B0;B1;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "a2")) then
- returnm ((vec_of_bits [B0;B1;B1;B0;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "a3")) then
- returnm ((vec_of_bits [B0;B1;B1;B0;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "a4")) then
- returnm ((vec_of_bits [B0;B1;B1;B1;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "a5")) then
- returnm ((vec_of_bits [B0;B1;B1;B1;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "a6")) then
- returnm ((vec_of_bits [B1;B0;B0;B0;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "a7")) then
- returnm ((vec_of_bits [B1;B0;B0;B0;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s2")) then
- returnm ((vec_of_bits [B1;B0;B0;B1;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s3")) then
- returnm ((vec_of_bits [B1;B0;B0;B1;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s4")) then
- returnm ((vec_of_bits [B1;B0;B1;B0;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s5")) then
- returnm ((vec_of_bits [B1;B0;B1;B0;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s6")) then
- returnm ((vec_of_bits [B1;B0;B1;B1;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s7")) then
- returnm ((vec_of_bits [B1;B0;B1;B1;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s8")) then
- returnm ((vec_of_bits [B1;B1;B0;B0;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s9")) then
- returnm ((vec_of_bits [B1;B1;B0;B0;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s10")) then
- returnm ((vec_of_bits [B1;B1;B0;B1;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "s11")) then
- returnm ((vec_of_bits [B1;B1;B0;B1;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "t3")) then
- returnm ((vec_of_bits [B1;B1;B1;B0;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "t4")) then
- returnm ((vec_of_bits [B1;B1;B1;B0;B1] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "t5")) then
- returnm ((vec_of_bits [B1;B1;B1;B1;B0] : mword 5)
- : mword 5)
- else if ((generic_eq p0_ "t6")) then
- returnm ((vec_of_bits [B1;B1;B1;B1;B1] : mword 5)
- : mword 5)
+ (if generic_eq p0_ "zero" then returnm ('b"00000" : mword 5)
+ else if generic_eq p0_ "ra" then returnm ('b"00001" : mword 5)
+ else if generic_eq p0_ "sp" then returnm ('b"00010" : mword 5)
+ else if generic_eq p0_ "gp" then returnm ('b"00011" : mword 5)
+ else if generic_eq p0_ "tp" then returnm ('b"00100" : mword 5)
+ else if generic_eq p0_ "t0" then returnm ('b"00101" : mword 5)
+ else if generic_eq p0_ "t1" then returnm ('b"00110" : mword 5)
+ else if generic_eq p0_ "t2" then returnm ('b"00111" : mword 5)
+ else if generic_eq p0_ "fp" then returnm ('b"01000" : mword 5)
+ else if generic_eq p0_ "s1" then returnm ('b"01001" : mword 5)
+ else if generic_eq p0_ "a0" then returnm ('b"01010" : mword 5)
+ else if generic_eq p0_ "a1" then returnm ('b"01011" : mword 5)
+ else if generic_eq p0_ "a2" then returnm ('b"01100" : mword 5)
+ else if generic_eq p0_ "a3" then returnm ('b"01101" : mword 5)
+ else if generic_eq p0_ "a4" then returnm ('b"01110" : mword 5)
+ else if generic_eq p0_ "a5" then returnm ('b"01111" : mword 5)
+ else if generic_eq p0_ "a6" then returnm ('b"10000" : mword 5)
+ else if generic_eq p0_ "a7" then returnm ('b"10001" : mword 5)
+ else if generic_eq p0_ "s2" then returnm ('b"10010" : mword 5)
+ else if generic_eq p0_ "s3" then returnm ('b"10011" : mword 5)
+ else if generic_eq p0_ "s4" then returnm ('b"10100" : mword 5)
+ else if generic_eq p0_ "s5" then returnm ('b"10101" : mword 5)
+ else if generic_eq p0_ "s6" then returnm ('b"10110" : mword 5)
+ else if generic_eq p0_ "s7" then returnm ('b"10111" : mword 5)
+ else if generic_eq p0_ "s8" then returnm ('b"11000" : mword 5)
+ else if generic_eq p0_ "s9" then returnm ('b"11001" : mword 5)
+ else if generic_eq p0_ "s10" then returnm ('b"11010" : mword 5)
+ else if generic_eq p0_ "s11" then returnm ('b"11011" : mword 5)
+ else if generic_eq p0_ "t3" then returnm ('b"11100" : mword 5)
+ else if generic_eq p0_ "t4" then returnm ('b"11101" : mword 5)
+ else if generic_eq p0_ "t5" then returnm ('b"11110" : mword 5)
+ else if generic_eq p0_ "t6" then returnm ('b"11111" : mword 5)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 5).
-Definition reg_name_forwards_matches (arg_ : mword 5)
-: bool :=
-
+Definition reg_name_forwards_matches (arg_ : mword 5) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B1;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B1;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B1;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B1;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B1;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B1;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B1] : mword 5))) then true
+ if eq_vec b__0 ('b"00000" : mword 5) then true
+ else if eq_vec b__0 ('b"00001" : mword 5) then true
+ else if eq_vec b__0 ('b"00010" : mword 5) then true
+ else if eq_vec b__0 ('b"00011" : mword 5) then true
+ else if eq_vec b__0 ('b"00100" : mword 5) then true
+ else if eq_vec b__0 ('b"00101" : mword 5) then true
+ else if eq_vec b__0 ('b"00110" : mword 5) then true
+ else if eq_vec b__0 ('b"00111" : mword 5) then true
+ else if eq_vec b__0 ('b"01000" : mword 5) then true
+ else if eq_vec b__0 ('b"01001" : mword 5) then true
+ else if eq_vec b__0 ('b"01010" : mword 5) then true
+ else if eq_vec b__0 ('b"01011" : mword 5) then true
+ else if eq_vec b__0 ('b"01100" : mword 5) then true
+ else if eq_vec b__0 ('b"01101" : mword 5) then true
+ else if eq_vec b__0 ('b"01110" : mword 5) then true
+ else if eq_vec b__0 ('b"01111" : mword 5) then true
+ else if eq_vec b__0 ('b"10000" : mword 5) then true
+ else if eq_vec b__0 ('b"10001" : mword 5) then true
+ else if eq_vec b__0 ('b"10010" : mword 5) then true
+ else if eq_vec b__0 ('b"10011" : mword 5) then true
+ else if eq_vec b__0 ('b"10100" : mword 5) then true
+ else if eq_vec b__0 ('b"10101" : mword 5) then true
+ else if eq_vec b__0 ('b"10110" : mword 5) then true
+ else if eq_vec b__0 ('b"10111" : mword 5) then true
+ else if eq_vec b__0 ('b"11000" : mword 5) then true
+ else if eq_vec b__0 ('b"11001" : mword 5) then true
+ else if eq_vec b__0 ('b"11010" : mword 5) then true
+ else if eq_vec b__0 ('b"11011" : mword 5) then true
+ else if eq_vec b__0 ('b"11100" : mword 5) then true
+ else if eq_vec b__0 ('b"11101" : mword 5) then true
+ else if eq_vec b__0 ('b"11110" : mword 5) then true
+ else if eq_vec b__0 ('b"11111" : mword 5) then true
else false.
-Definition reg_name_backwards_matches (arg_ : string)
-: bool :=
-
+Definition reg_name_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "zero")) then true
- else if ((generic_eq p0_ "ra")) then true
- else if ((generic_eq p0_ "sp")) then true
- else if ((generic_eq p0_ "gp")) then true
- else if ((generic_eq p0_ "tp")) then true
- else if ((generic_eq p0_ "t0")) then true
- else if ((generic_eq p0_ "t1")) then true
- else if ((generic_eq p0_ "t2")) then true
- else if ((generic_eq p0_ "fp")) then true
- else if ((generic_eq p0_ "s1")) then true
- else if ((generic_eq p0_ "a0")) then true
- else if ((generic_eq p0_ "a1")) then true
- else if ((generic_eq p0_ "a2")) then true
- else if ((generic_eq p0_ "a3")) then true
- else if ((generic_eq p0_ "a4")) then true
- else if ((generic_eq p0_ "a5")) then true
- else if ((generic_eq p0_ "a6")) then true
- else if ((generic_eq p0_ "a7")) then true
- else if ((generic_eq p0_ "s2")) then true
- else if ((generic_eq p0_ "s3")) then true
- else if ((generic_eq p0_ "s4")) then true
- else if ((generic_eq p0_ "s5")) then true
- else if ((generic_eq p0_ "s6")) then true
- else if ((generic_eq p0_ "s7")) then true
- else if ((generic_eq p0_ "s8")) then true
- else if ((generic_eq p0_ "s9")) then true
- else if ((generic_eq p0_ "s10")) then true
- else if ((generic_eq p0_ "s11")) then true
- else if ((generic_eq p0_ "t3")) then true
- else if ((generic_eq p0_ "t4")) then true
- else if ((generic_eq p0_ "t5")) then true
- else if ((generic_eq p0_ "t6")) then true
+ if generic_eq p0_ "zero" then true
+ else if generic_eq p0_ "ra" then true
+ else if generic_eq p0_ "sp" then true
+ else if generic_eq p0_ "gp" then true
+ else if generic_eq p0_ "tp" then true
+ else if generic_eq p0_ "t0" then true
+ else if generic_eq p0_ "t1" then true
+ else if generic_eq p0_ "t2" then true
+ else if generic_eq p0_ "fp" then true
+ else if generic_eq p0_ "s1" then true
+ else if generic_eq p0_ "a0" then true
+ else if generic_eq p0_ "a1" then true
+ else if generic_eq p0_ "a2" then true
+ else if generic_eq p0_ "a3" then true
+ else if generic_eq p0_ "a4" then true
+ else if generic_eq p0_ "a5" then true
+ else if generic_eq p0_ "a6" then true
+ else if generic_eq p0_ "a7" then true
+ else if generic_eq p0_ "s2" then true
+ else if generic_eq p0_ "s3" then true
+ else if generic_eq p0_ "s4" then true
+ else if generic_eq p0_ "s5" then true
+ else if generic_eq p0_ "s6" then true
+ else if generic_eq p0_ "s7" then true
+ else if generic_eq p0_ "s8" then true
+ else if generic_eq p0_ "s9" then true
+ else if generic_eq p0_ "s10" then true
+ else if generic_eq p0_ "s11" then true
+ else if generic_eq p0_ "t3" then true
+ else if generic_eq p0_ "t4" then true
+ else if generic_eq p0_ "t5" then true
+ else if generic_eq p0_ "t6" then true
else false.
-Definition _s164_ (_s165_ : string)
-: option string :=
-
+Definition _s164_ (_s165_ : string) : option string :=
let _s166_ := _s165_ in
- if ((string_startswith _s166_ "t6")) then
- match (string_drop _s166_ (projT1 (string_length "t6"))) with | s_ => Some (s_) end
+ if string_startswith _s166_ "t6" then
+ match (string_drop _s166_ (projT1 (string_length "t6"))) with | s_ => Some s_ end
else None.
-Definition _s160_ (_s161_ : string)
-: option string :=
-
+Definition _s160_ (_s161_ : string) : option string :=
let _s162_ := _s161_ in
- if ((string_startswith _s162_ "t5")) then
- match (string_drop _s162_ (projT1 (string_length "t5"))) with | s_ => Some (s_) end
+ if string_startswith _s162_ "t5" then
+ match (string_drop _s162_ (projT1 (string_length "t5"))) with | s_ => Some s_ end
else None.
-Definition _s156_ (_s157_ : string)
-: option string :=
-
+Definition _s156_ (_s157_ : string) : option string :=
let _s158_ := _s157_ in
- if ((string_startswith _s158_ "t4")) then
- match (string_drop _s158_ (projT1 (string_length "t4"))) with | s_ => Some (s_) end
+ if string_startswith _s158_ "t4" then
+ match (string_drop _s158_ (projT1 (string_length "t4"))) with | s_ => Some s_ end
else None.
-Definition _s152_ (_s153_ : string)
-: option string :=
-
+Definition _s152_ (_s153_ : string) : option string :=
let _s154_ := _s153_ in
- if ((string_startswith _s154_ "t3")) then
- match (string_drop _s154_ (projT1 (string_length "t3"))) with | s_ => Some (s_) end
+ if string_startswith _s154_ "t3" then
+ match (string_drop _s154_ (projT1 (string_length "t3"))) with | s_ => Some s_ end
else None.
-Definition _s148_ (_s149_ : string)
-: option string :=
-
+Definition _s148_ (_s149_ : string) : option string :=
let _s150_ := _s149_ in
- if ((string_startswith _s150_ "s11")) then
- match (string_drop _s150_ (projT1 (string_length "s11"))) with | s_ => Some (s_) end
+ if string_startswith _s150_ "s11" then
+ match (string_drop _s150_ (projT1 (string_length "s11"))) with | s_ => Some s_ end
else None.
-Definition _s144_ (_s145_ : string)
-: option string :=
-
+Definition _s144_ (_s145_ : string) : option string :=
let _s146_ := _s145_ in
- if ((string_startswith _s146_ "s10")) then
- match (string_drop _s146_ (projT1 (string_length "s10"))) with | s_ => Some (s_) end
+ if string_startswith _s146_ "s10" then
+ match (string_drop _s146_ (projT1 (string_length "s10"))) with | s_ => Some s_ end
else None.
-Definition _s140_ (_s141_ : string)
-: option string :=
-
+Definition _s140_ (_s141_ : string) : option string :=
let _s142_ := _s141_ in
- if ((string_startswith _s142_ "s9")) then
- match (string_drop _s142_ (projT1 (string_length "s9"))) with | s_ => Some (s_) end
+ if string_startswith _s142_ "s9" then
+ match (string_drop _s142_ (projT1 (string_length "s9"))) with | s_ => Some s_ end
else None.
-Definition _s136_ (_s137_ : string)
-: option string :=
-
+Definition _s136_ (_s137_ : string) : option string :=
let _s138_ := _s137_ in
- if ((string_startswith _s138_ "s8")) then
- match (string_drop _s138_ (projT1 (string_length "s8"))) with | s_ => Some (s_) end
+ if string_startswith _s138_ "s8" then
+ match (string_drop _s138_ (projT1 (string_length "s8"))) with | s_ => Some s_ end
else None.
-Definition _s132_ (_s133_ : string)
-: option string :=
-
+Definition _s132_ (_s133_ : string) : option string :=
let _s134_ := _s133_ in
- if ((string_startswith _s134_ "s7")) then
- match (string_drop _s134_ (projT1 (string_length "s7"))) with | s_ => Some (s_) end
+ if string_startswith _s134_ "s7" then
+ match (string_drop _s134_ (projT1 (string_length "s7"))) with | s_ => Some s_ end
else None.
-Definition _s128_ (_s129_ : string)
-: option string :=
-
+Definition _s128_ (_s129_ : string) : option string :=
let _s130_ := _s129_ in
- if ((string_startswith _s130_ "s6")) then
- match (string_drop _s130_ (projT1 (string_length "s6"))) with | s_ => Some (s_) end
+ if string_startswith _s130_ "s6" then
+ match (string_drop _s130_ (projT1 (string_length "s6"))) with | s_ => Some s_ end
else None.
-Definition _s124_ (_s125_ : string)
-: option string :=
-
+Definition _s124_ (_s125_ : string) : option string :=
let _s126_ := _s125_ in
- if ((string_startswith _s126_ "s5")) then
- match (string_drop _s126_ (projT1 (string_length "s5"))) with | s_ => Some (s_) end
+ if string_startswith _s126_ "s5" then
+ match (string_drop _s126_ (projT1 (string_length "s5"))) with | s_ => Some s_ end
else None.
-Definition _s120_ (_s121_ : string)
-: option string :=
-
+Definition _s120_ (_s121_ : string) : option string :=
let _s122_ := _s121_ in
- if ((string_startswith _s122_ "s4")) then
- match (string_drop _s122_ (projT1 (string_length "s4"))) with | s_ => Some (s_) end
+ if string_startswith _s122_ "s4" then
+ match (string_drop _s122_ (projT1 (string_length "s4"))) with | s_ => Some s_ end
else None.
-Definition _s116_ (_s117_ : string)
-: option string :=
-
+Definition _s116_ (_s117_ : string) : option string :=
let _s118_ := _s117_ in
- if ((string_startswith _s118_ "s3")) then
- match (string_drop _s118_ (projT1 (string_length "s3"))) with | s_ => Some (s_) end
+ if string_startswith _s118_ "s3" then
+ match (string_drop _s118_ (projT1 (string_length "s3"))) with | s_ => Some s_ end
else None.
-Definition _s112_ (_s113_ : string)
-: option string :=
-
+Definition _s112_ (_s113_ : string) : option string :=
let _s114_ := _s113_ in
- if ((string_startswith _s114_ "s2")) then
- match (string_drop _s114_ (projT1 (string_length "s2"))) with | s_ => Some (s_) end
+ if string_startswith _s114_ "s2" then
+ match (string_drop _s114_ (projT1 (string_length "s2"))) with | s_ => Some s_ end
else None.
-Definition _s108_ (_s109_ : string)
-: option string :=
-
+Definition _s108_ (_s109_ : string) : option string :=
let _s110_ := _s109_ in
- if ((string_startswith _s110_ "a7")) then
- match (string_drop _s110_ (projT1 (string_length "a7"))) with | s_ => Some (s_) end
+ if string_startswith _s110_ "a7" then
+ match (string_drop _s110_ (projT1 (string_length "a7"))) with | s_ => Some s_ end
else None.
-Definition _s104_ (_s105_ : string)
-: option string :=
-
+Definition _s104_ (_s105_ : string) : option string :=
let _s106_ := _s105_ in
- if ((string_startswith _s106_ "a6")) then
- match (string_drop _s106_ (projT1 (string_length "a6"))) with | s_ => Some (s_) end
+ if string_startswith _s106_ "a6" then
+ match (string_drop _s106_ (projT1 (string_length "a6"))) with | s_ => Some s_ end
else None.
-Definition _s100_ (_s101_ : string)
-: option string :=
-
+Definition _s100_ (_s101_ : string) : option string :=
let _s102_ := _s101_ in
- if ((string_startswith _s102_ "a5")) then
- match (string_drop _s102_ (projT1 (string_length "a5"))) with | s_ => Some (s_) end
+ if string_startswith _s102_ "a5" then
+ match (string_drop _s102_ (projT1 (string_length "a5"))) with | s_ => Some s_ end
else None.
-Definition _s96_ (_s97_ : string)
-: option string :=
-
+Definition _s96_ (_s97_ : string) : option string :=
let _s98_ := _s97_ in
- if ((string_startswith _s98_ "a4")) then
- match (string_drop _s98_ (projT1 (string_length "a4"))) with | s_ => Some (s_) end
+ if string_startswith _s98_ "a4" then
+ match (string_drop _s98_ (projT1 (string_length "a4"))) with | s_ => Some s_ end
else None.
-Definition _s92_ (_s93_ : string)
-: option string :=
-
+Definition _s92_ (_s93_ : string) : option string :=
let _s94_ := _s93_ in
- if ((string_startswith _s94_ "a3")) then
- match (string_drop _s94_ (projT1 (string_length "a3"))) with | s_ => Some (s_) end
+ if string_startswith _s94_ "a3" then
+ match (string_drop _s94_ (projT1 (string_length "a3"))) with | s_ => Some s_ end
else None.
-Definition _s88_ (_s89_ : string)
-: option string :=
-
+Definition _s88_ (_s89_ : string) : option string :=
let _s90_ := _s89_ in
- if ((string_startswith _s90_ "a2")) then
- match (string_drop _s90_ (projT1 (string_length "a2"))) with | s_ => Some (s_) end
+ if string_startswith _s90_ "a2" then
+ match (string_drop _s90_ (projT1 (string_length "a2"))) with | s_ => Some s_ end
else None.
-Definition _s84_ (_s85_ : string)
-: option string :=
-
+Definition _s84_ (_s85_ : string) : option string :=
let _s86_ := _s85_ in
- if ((string_startswith _s86_ "a1")) then
- match (string_drop _s86_ (projT1 (string_length "a1"))) with | s_ => Some (s_) end
+ if string_startswith _s86_ "a1" then
+ match (string_drop _s86_ (projT1 (string_length "a1"))) with | s_ => Some s_ end
else None.
-Definition _s80_ (_s81_ : string)
-: option string :=
-
+Definition _s80_ (_s81_ : string) : option string :=
let _s82_ := _s81_ in
- if ((string_startswith _s82_ "a0")) then
- match (string_drop _s82_ (projT1 (string_length "a0"))) with | s_ => Some (s_) end
+ if string_startswith _s82_ "a0" then
+ match (string_drop _s82_ (projT1 (string_length "a0"))) with | s_ => Some s_ end
else None.
-Definition _s76_ (_s77_ : string)
-: option string :=
-
+Definition _s76_ (_s77_ : string) : option string :=
let _s78_ := _s77_ in
- if ((string_startswith _s78_ "s1")) then
- match (string_drop _s78_ (projT1 (string_length "s1"))) with | s_ => Some (s_) end
+ if string_startswith _s78_ "s1" then
+ match (string_drop _s78_ (projT1 (string_length "s1"))) with | s_ => Some s_ end
else None.
-Definition _s72_ (_s73_ : string)
-: option string :=
-
+Definition _s72_ (_s73_ : string) : option string :=
let _s74_ := _s73_ in
- if ((string_startswith _s74_ "fp")) then
- match (string_drop _s74_ (projT1 (string_length "fp"))) with | s_ => Some (s_) end
+ if string_startswith _s74_ "fp" then
+ match (string_drop _s74_ (projT1 (string_length "fp"))) with | s_ => Some s_ end
else None.
-Definition _s68_ (_s69_ : string)
-: option string :=
-
+Definition _s68_ (_s69_ : string) : option string :=
let _s70_ := _s69_ in
- if ((string_startswith _s70_ "t2")) then
- match (string_drop _s70_ (projT1 (string_length "t2"))) with | s_ => Some (s_) end
+ if string_startswith _s70_ "t2" then
+ match (string_drop _s70_ (projT1 (string_length "t2"))) with | s_ => Some s_ end
else None.
-Definition _s64_ (_s65_ : string)
-: option string :=
-
+Definition _s64_ (_s65_ : string) : option string :=
let _s66_ := _s65_ in
- if ((string_startswith _s66_ "t1")) then
- match (string_drop _s66_ (projT1 (string_length "t1"))) with | s_ => Some (s_) end
+ if string_startswith _s66_ "t1" then
+ match (string_drop _s66_ (projT1 (string_length "t1"))) with | s_ => Some s_ end
else None.
-Definition _s60_ (_s61_ : string)
-: option string :=
-
+Definition _s60_ (_s61_ : string) : option string :=
let _s62_ := _s61_ in
- if ((string_startswith _s62_ "t0")) then
- match (string_drop _s62_ (projT1 (string_length "t0"))) with | s_ => Some (s_) end
+ if string_startswith _s62_ "t0" then
+ match (string_drop _s62_ (projT1 (string_length "t0"))) with | s_ => Some s_ end
else None.
-Definition _s56_ (_s57_ : string)
-: option string :=
-
+Definition _s56_ (_s57_ : string) : option string :=
let _s58_ := _s57_ in
- if ((string_startswith _s58_ "tp")) then
- match (string_drop _s58_ (projT1 (string_length "tp"))) with | s_ => Some (s_) end
+ if string_startswith _s58_ "tp" then
+ match (string_drop _s58_ (projT1 (string_length "tp"))) with | s_ => Some s_ end
else None.
-Definition _s52_ (_s53_ : string)
-: option string :=
-
+Definition _s52_ (_s53_ : string) : option string :=
let _s54_ := _s53_ in
- if ((string_startswith _s54_ "gp")) then
- match (string_drop _s54_ (projT1 (string_length "gp"))) with | s_ => Some (s_) end
+ if string_startswith _s54_ "gp" then
+ match (string_drop _s54_ (projT1 (string_length "gp"))) with | s_ => Some s_ end
else None.
-Definition _s48_ (_s49_ : string)
-: option string :=
-
+Definition _s48_ (_s49_ : string) : option string :=
let _s50_ := _s49_ in
- if ((string_startswith _s50_ "sp")) then
- match (string_drop _s50_ (projT1 (string_length "sp"))) with | s_ => Some (s_) end
+ if string_startswith _s50_ "sp" then
+ match (string_drop _s50_ (projT1 (string_length "sp"))) with | s_ => Some s_ end
else None.
-Definition _s44_ (_s45_ : string)
-: option string :=
-
+Definition _s44_ (_s45_ : string) : option string :=
let _s46_ := _s45_ in
- if ((string_startswith _s46_ "ra")) then
- match (string_drop _s46_ (projT1 (string_length "ra"))) with | s_ => Some (s_) end
+ if string_startswith _s46_ "ra" then
+ match (string_drop _s46_ (projT1 (string_length "ra"))) with | s_ => Some s_ end
else None.
-Definition _s40_ (_s41_ : string)
-: option string :=
-
+Definition _s40_ (_s41_ : string) : option string :=
let _s42_ := _s41_ in
- if ((string_startswith _s42_ "zero")) then
- match (string_drop _s42_ (projT1 (string_length "zero"))) with | s_ => Some (s_) end
+ if string_startswith _s42_ "zero" then
+ match (string_drop _s42_ (projT1 (string_length "zero"))) with | s_ => Some s_ end
else None.
-Definition reg_name_matches_prefix (arg_ : string)
-: M (option ((mword 5 * {n : Z & ArithFact (n >= 0)}))) :=
-
+Definition reg_name_matches_prefix (arg_ : string)
+: M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)}))) :=
let _s43_ := arg_ in
- (if ((match (_s40_ _s43_) with | Some (s_) => true | _ => false end)) then
+ (if match (_s40_ _s43_) with | Some s_ => true | _ => false end then
(match (_s40_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s44_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s44_ _s43_) with | Some s_ => true | _ => false end then
(match (_s44_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s48_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s48_ _s43_) with | Some s_ => true | _ => false end then
(match (_s48_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s52_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s52_ _s43_) with | Some s_ => true | _ => false end then
(match (_s52_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s56_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s56_ _s43_) with | Some s_ => true | _ => false end then
(match (_s56_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B0;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s60_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s60_ _s43_) with | Some s_ => true | _ => false end then
(match (_s60_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B0;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s64_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s64_ _s43_) with | Some s_ => true | _ => false end then
(match (_s64_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s68_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s68_ _s43_) with | Some s_ => true | _ => false end then
(match (_s68_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s72_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s72_ _s43_) with | Some s_ => true | _ => false end then
(match (_s72_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B0;B0;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s76_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s76_ _s43_) with | Some s_ => true | _ => false end then
(match (_s76_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B0;B0;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s80_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s80_ _s43_) with | Some s_ => true | _ => false end then
(match (_s80_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B0;B1;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s84_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s84_ _s43_) with | Some s_ => true | _ => false end then
(match (_s84_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B0;B1;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s88_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s88_ _s43_) with | Some s_ => true | _ => false end then
(match (_s88_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1;B0;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s92_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s92_ _s43_) with | Some s_ => true | _ => false end then
(match (_s92_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1;B0;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s96_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s96_ _s43_) with | Some s_ => true | _ => false end then
(match (_s96_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1;B1;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s100_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s100_ _s43_) with | Some s_ => true | _ => false end then
(match (_s100_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1;B1;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s104_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s104_ _s43_) with | Some s_ => true | _ => false end then
(match (_s104_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B0;B0;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s108_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s108_ _s43_) with | Some s_ => true | _ => false end then
(match (_s108_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B0;B0;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s112_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s112_ _s43_) with | Some s_ => true | _ => false end then
(match (_s112_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B0;B1;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s116_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s116_ _s43_) with | Some s_ => true | _ => false end then
(match (_s116_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B0;B1;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s120_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s120_ _s43_) with | Some s_ => true | _ => false end then
(match (_s120_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1;B0;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s124_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s124_ _s43_) with | Some s_ => true | _ => false end then
(match (_s124_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1;B0;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s128_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s128_ _s43_) with | Some s_ => true | _ => false end then
(match (_s128_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1;B1;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s132_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s132_ _s43_) with | Some s_ => true | _ => false end then
(match (_s132_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1;B1;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s136_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s136_ _s43_) with | Some s_ => true | _ => false end then
(match (_s136_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B0;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s140_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s140_ _s43_) with | Some s_ => true | _ => false end then
(match (_s140_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B0;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s144_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s144_ _s43_) with | Some s_ => true | _ => false end then
(match (_s144_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B1;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s148_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s148_ _s43_) with | Some s_ => true | _ => false end then
(match (_s148_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B1;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s152_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s152_ _s43_) with | Some s_ => true | _ => false end then
(match (_s152_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1;B0;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s156_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s156_ _s43_) with | Some s_ => true | _ => false end then
(match (_s156_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1;B0;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s160_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s160_ _s43_) with | Some s_ => true | _ => false end then
(match (_s160_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1;B1;B0] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s164_ _s43_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s164_ _s43_) with | Some s_ => true | _ => false end then
(match (_s164_ _s43_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1;B1;B1] : mword 5), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((mword 5 * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((mword 5 * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)}))).
-Definition creg_name_forwards (arg_ : mword 3)
-: M (string) :=
-
+Definition creg_name_forwards (arg_ : mword 3) : M (string) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then returnm ("s0" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then returnm ("s1" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then returnm ("a0" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then returnm ("a1" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then returnm ("a2" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then returnm ("a3" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then returnm ("a4" : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1] : mword 3))) then returnm ("a5" : string)
+ (if eq_vec b__0 ('b"000" : mword 3) then returnm "s0"
+ else if eq_vec b__0 ('b"001" : mword 3) then returnm "s1"
+ else if eq_vec b__0 ('b"010" : mword 3) then returnm "a0"
+ else if eq_vec b__0 ('b"011" : mword 3) then returnm "a1"
+ else if eq_vec b__0 ('b"100" : mword 3) then returnm "a2"
+ else if eq_vec b__0 ('b"101" : mword 3) then returnm "a3"
+ else if eq_vec b__0 ('b"110" : mword 3) then returnm "a4"
+ else if eq_vec b__0 ('b"111" : mword 3) then returnm "a5"
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string).
-Definition creg_name_backwards (arg_ : string)
-: M (mword 3) :=
-
+Definition creg_name_backwards (arg_ : string) : M (mword 3) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "s0")) then returnm ((vec_of_bits [B0;B0;B0] : mword 3) : mword 3)
- else if ((generic_eq p0_ "s1")) then returnm ((vec_of_bits [B0;B0;B1] : mword 3) : mword 3)
- else if ((generic_eq p0_ "a0")) then returnm ((vec_of_bits [B0;B1;B0] : mword 3) : mword 3)
- else if ((generic_eq p0_ "a1")) then returnm ((vec_of_bits [B0;B1;B1] : mword 3) : mword 3)
- else if ((generic_eq p0_ "a2")) then returnm ((vec_of_bits [B1;B0;B0] : mword 3) : mword 3)
- else if ((generic_eq p0_ "a3")) then returnm ((vec_of_bits [B1;B0;B1] : mword 3) : mword 3)
- else if ((generic_eq p0_ "a4")) then returnm ((vec_of_bits [B1;B1;B0] : mword 3) : mword 3)
- else if ((generic_eq p0_ "a5")) then returnm ((vec_of_bits [B1;B1;B1] : mword 3) : mword 3)
+ (if generic_eq p0_ "s0" then returnm ('b"000" : mword 3)
+ else if generic_eq p0_ "s1" then returnm ('b"001" : mword 3)
+ else if generic_eq p0_ "a0" then returnm ('b"010" : mword 3)
+ else if generic_eq p0_ "a1" then returnm ('b"011" : mword 3)
+ else if generic_eq p0_ "a2" then returnm ('b"100" : mword 3)
+ else if generic_eq p0_ "a3" then returnm ('b"101" : mword 3)
+ else if generic_eq p0_ "a4" then returnm ('b"110" : mword 3)
+ else if generic_eq p0_ "a5" then returnm ('b"111" : mword 3)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 3).
-Definition creg_name_forwards_matches (arg_ : mword 3)
-: bool :=
-
+Definition creg_name_forwards_matches (arg_ : mword 3) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1] : mword 3))) then true
+ if eq_vec b__0 ('b"000" : mword 3) then true
+ else if eq_vec b__0 ('b"001" : mword 3) then true
+ else if eq_vec b__0 ('b"010" : mword 3) then true
+ else if eq_vec b__0 ('b"011" : mword 3) then true
+ else if eq_vec b__0 ('b"100" : mword 3) then true
+ else if eq_vec b__0 ('b"101" : mword 3) then true
+ else if eq_vec b__0 ('b"110" : mword 3) then true
+ else if eq_vec b__0 ('b"111" : mword 3) then true
else false.
-Definition creg_name_backwards_matches (arg_ : string)
-: bool :=
-
+Definition creg_name_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "s0")) then true
- else if ((generic_eq p0_ "s1")) then true
- else if ((generic_eq p0_ "a0")) then true
- else if ((generic_eq p0_ "a1")) then true
- else if ((generic_eq p0_ "a2")) then true
- else if ((generic_eq p0_ "a3")) then true
- else if ((generic_eq p0_ "a4")) then true
- else if ((generic_eq p0_ "a5")) then true
+ if generic_eq p0_ "s0" then true
+ else if generic_eq p0_ "s1" then true
+ else if generic_eq p0_ "a0" then true
+ else if generic_eq p0_ "a1" then true
+ else if generic_eq p0_ "a2" then true
+ else if generic_eq p0_ "a3" then true
+ else if generic_eq p0_ "a4" then true
+ else if generic_eq p0_ "a5" then true
else false.
-Definition _s196_ (_s197_ : string)
-: option string :=
-
+Definition _s196_ (_s197_ : string) : option string :=
let _s198_ := _s197_ in
- if ((string_startswith _s198_ "a5")) then
- match (string_drop _s198_ (projT1 (string_length "a5"))) with | s_ => Some (s_) end
+ if string_startswith _s198_ "a5" then
+ match (string_drop _s198_ (projT1 (string_length "a5"))) with | s_ => Some s_ end
else None.
-Definition _s192_ (_s193_ : string)
-: option string :=
-
+Definition _s192_ (_s193_ : string) : option string :=
let _s194_ := _s193_ in
- if ((string_startswith _s194_ "a4")) then
- match (string_drop _s194_ (projT1 (string_length "a4"))) with | s_ => Some (s_) end
+ if string_startswith _s194_ "a4" then
+ match (string_drop _s194_ (projT1 (string_length "a4"))) with | s_ => Some s_ end
else None.
-Definition _s188_ (_s189_ : string)
-: option string :=
-
+Definition _s188_ (_s189_ : string) : option string :=
let _s190_ := _s189_ in
- if ((string_startswith _s190_ "a3")) then
- match (string_drop _s190_ (projT1 (string_length "a3"))) with | s_ => Some (s_) end
+ if string_startswith _s190_ "a3" then
+ match (string_drop _s190_ (projT1 (string_length "a3"))) with | s_ => Some s_ end
else None.
-Definition _s184_ (_s185_ : string)
-: option string :=
-
+Definition _s184_ (_s185_ : string) : option string :=
let _s186_ := _s185_ in
- if ((string_startswith _s186_ "a2")) then
- match (string_drop _s186_ (projT1 (string_length "a2"))) with | s_ => Some (s_) end
+ if string_startswith _s186_ "a2" then
+ match (string_drop _s186_ (projT1 (string_length "a2"))) with | s_ => Some s_ end
else None.
-Definition _s180_ (_s181_ : string)
-: option string :=
-
+Definition _s180_ (_s181_ : string) : option string :=
let _s182_ := _s181_ in
- if ((string_startswith _s182_ "a1")) then
- match (string_drop _s182_ (projT1 (string_length "a1"))) with | s_ => Some (s_) end
+ if string_startswith _s182_ "a1" then
+ match (string_drop _s182_ (projT1 (string_length "a1"))) with | s_ => Some s_ end
else None.
-Definition _s176_ (_s177_ : string)
-: option string :=
-
+Definition _s176_ (_s177_ : string) : option string :=
let _s178_ := _s177_ in
- if ((string_startswith _s178_ "a0")) then
- match (string_drop _s178_ (projT1 (string_length "a0"))) with | s_ => Some (s_) end
+ if string_startswith _s178_ "a0" then
+ match (string_drop _s178_ (projT1 (string_length "a0"))) with | s_ => Some s_ end
else None.
-Definition _s172_ (_s173_ : string)
-: option string :=
-
+Definition _s172_ (_s173_ : string) : option string :=
let _s174_ := _s173_ in
- if ((string_startswith _s174_ "s1")) then
- match (string_drop _s174_ (projT1 (string_length "s1"))) with | s_ => Some (s_) end
+ if string_startswith _s174_ "s1" then
+ match (string_drop _s174_ (projT1 (string_length "s1"))) with | s_ => Some s_ end
else None.
-Definition _s168_ (_s169_ : string)
-: option string :=
-
+Definition _s168_ (_s169_ : string) : option string :=
let _s170_ := _s169_ in
- if ((string_startswith _s170_ "s0")) then
- match (string_drop _s170_ (projT1 (string_length "s0"))) with | s_ => Some (s_) end
+ if string_startswith _s170_ "s0" then
+ match (string_drop _s170_ (projT1 (string_length "s0"))) with | s_ => Some s_ end
else None.
-Definition creg_name_matches_prefix (arg_ : string)
-: M (option ((mword 3 * {n : Z & ArithFact (n >= 0)}))) :=
-
+Definition creg_name_matches_prefix (arg_ : string)
+: M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)}))) :=
let _s171_ := arg_ in
- (if ((match (_s168_ _s171_) with | Some (s_) => true | _ => false end)) then
+ (if match (_s168_ _s171_) with | Some s_ => true | _ => false end then
(match (_s168_ _s171_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0] : mword 3), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"000"
+ : mword 3, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s172_ _s171_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s172_ _s171_) with | Some s_ => true | _ => false end then
(match (_s172_ _s171_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1] : mword 3), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"001"
+ : mword 3, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s176_ _s171_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s176_ _s171_) with | Some s_ => true | _ => false end then
(match (_s176_ _s171_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B0] : mword 3), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"010"
+ : mword 3, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s180_ _s171_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s180_ _s171_) with | Some s_ => true | _ => false end then
(match (_s180_ _s171_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1] : mword 3), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"011"
+ : mword 3, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s184_ _s171_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s184_ _s171_) with | Some s_ => true | _ => false end then
(match (_s184_ _s171_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B0] : mword 3), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"100"
+ : mword 3, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s188_ _s171_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s188_ _s171_) with | Some s_ => true | _ => false end then
(match (_s188_ _s171_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1] : mword 3), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"101"
+ : mword 3, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s192_ _s171_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s192_ _s171_) with | Some s_ => true | _ => false end then
(match (_s192_ _s171_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0] : mword 3), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"110"
+ : mword 3, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s196_ _s171_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s196_ _s171_) with | Some s_ => true | _ => false end then
(match (_s196_ _s171_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1] : mword 3), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"111"
+ : mword 3, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((mword 3 * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((mword 3 * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 3 * {n : Z & ArithFact (n >=? 0)}))).
-Definition init_base_regs '(tt : unit)
-: M (unit) :=
-
+Definition init_base_regs '(tt : unit) : M (unit) :=
write_reg x1_ref zero_reg >>
write_reg x2_ref zero_reg >>
write_reg x3_ref zero_reg >>
@@ -3740,34 +3394,25 @@ Definition init_base_regs '(tt : unit)
write_reg x29_ref zero_reg >>
write_reg x30_ref zero_reg >> write_reg x31_ref zero_reg : M (unit).
-Definition get_arch_pc '(tt : unit)
-: M (mword 64) :=
-
- ((read_reg PC_ref) : M (mword 64))
- : M (mword 64).
+Definition get_arch_pc '(tt : unit) : M (mword 64) :=
+ ((read_reg PC_ref) : M (mword 64)) : M (mword 64).
-Definition get_next_pc '(tt : unit)
-: M (mword 64) :=
-
- ((read_reg nextPC_ref) : M (mword 64))
- : M (mword 64).
+Definition get_next_pc '(tt : unit) : M (mword 64) :=
+ ((read_reg nextPC_ref) : M (mword 64)) : M (mword 64).
-Definition set_next_pc (pc : mword 64) : M (unit) := write_reg nextPC_ref pc : M (unit).
+Definition set_next_pc (pc : mword 64) : M (unit) := write_reg nextPC_ref pc : M (unit).
-Definition tick_pc '(tt : unit)
-: M (unit) :=
-
+Definition tick_pc '(tt : unit) : M (unit) :=
((read_reg nextPC_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
write_reg PC_ref w__0
: M (unit).
-Definition Mk_Misa (v : mword 64) : Misa := {| Misa_Misa_chunk_0 := (subrange_vec_dec v 63 0) |}.
+Definition Mk_Misa (v : mword 64) : Misa := {| Misa_Misa_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Misa_bits (v : Misa) : mword 64 := subrange_vec_dec v.(Misa_Misa_chunk_0) 63 0.
+Definition _get_Misa_bits (v : Misa) : mword 64 := subrange_vec_dec v.(Misa_Misa_chunk_0) 63 0.
-Definition _set_Misa_bits (r_ref : register_ref regstate register_value Misa) (v : mword 64)
+Definition _set_Misa_bits (r_ref : register_ref regstate register_value Misa) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3777,18 +3422,15 @@ Definition _set_Misa_bits (r_ref : register_ref regstate register_value Misa) (v
write_reg r_ref r
: M (unit).
-Definition _update_Misa_bits (v : Misa) (x : mword 64)
-: Misa :=
-
+Definition _update_Misa_bits (v : Misa) (x : mword 64) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Misa_MXL (v : Misa) : mword 2 := subrange_vec_dec v.(Misa_Misa_chunk_0) 63 62.
+Definition _get_Misa_MXL (v : Misa) : mword 2 := subrange_vec_dec v.(Misa_Misa_chunk_0) 63 62.
-Definition _set_Misa_MXL (r_ref : register_ref regstate register_value Misa) (v : mword 2)
+Definition _set_Misa_MXL (r_ref : register_ref regstate register_value Misa) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3798,18 +3440,14 @@ Definition _set_Misa_MXL (r_ref : register_ref regstate register_value Misa) (v
write_reg r_ref r
: M (unit).
-Definition _update_Misa_MXL (v : Misa) (x : mword 2)
-: Misa :=
-
+Definition _update_Misa_MXL (v : Misa) (x : mword 2) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 63 62 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_Misa_Z (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 25 25.
+Definition _get_Misa_Z (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 25 25.
-Definition _set_Misa_Z (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_Z (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3819,18 +3457,14 @@ Definition _set_Misa_Z (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_Z (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_Z (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 25 25 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_Y (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 24 24.
+Definition _get_Misa_Y (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 24 24.
-Definition _set_Misa_Y (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_Y (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3840,18 +3474,14 @@ Definition _set_Misa_Y (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_Y (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_Y (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 24 24 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_X (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 23 23.
+Definition _get_Misa_X (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 23 23.
-Definition _set_Misa_X (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_X (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3861,18 +3491,14 @@ Definition _set_Misa_X (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_X (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_X (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 23 23 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_W (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 22 22.
+Definition _get_Misa_W (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 22 22.
-Definition _set_Misa_W (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_W (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3882,18 +3508,14 @@ Definition _set_Misa_W (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_W (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_W (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 22 22 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_V (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 21 21.
+Definition _get_Misa_V (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 21 21.
-Definition _set_Misa_V (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_V (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3903,18 +3525,14 @@ Definition _set_Misa_V (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_V (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_V (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 21 21 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_U (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 20 20.
+Definition _get_Misa_U (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 20 20.
-Definition _set_Misa_U (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_U (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3924,18 +3542,14 @@ Definition _set_Misa_U (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_U (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_U (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 20 20 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_T (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 19 19.
+Definition _get_Misa_T (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 19 19.
-Definition _set_Misa_T (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_T (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3945,18 +3559,14 @@ Definition _set_Misa_T (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_T (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_T (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 19 19 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_S (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 18 18.
+Definition _get_Misa_S (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 18 18.
-Definition _set_Misa_S (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_S (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3966,18 +3576,14 @@ Definition _set_Misa_S (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_S (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_S (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 18 18 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_R (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 17 17.
+Definition _get_Misa_R (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 17 17.
-Definition _set_Misa_R (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_R (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -3987,18 +3593,14 @@ Definition _set_Misa_R (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_R (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_R (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 17 17 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_Q (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 16 16.
+Definition _get_Misa_Q (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 16 16.
-Definition _set_Misa_Q (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_Q (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4008,18 +3610,14 @@ Definition _set_Misa_Q (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_Q (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_Q (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 16 16 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_P (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 15 15.
+Definition _get_Misa_P (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 15 15.
-Definition _set_Misa_P (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_P (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4029,18 +3627,14 @@ Definition _set_Misa_P (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_P (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_P (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 15 15 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_O (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 14 14.
+Definition _get_Misa_O (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 14 14.
-Definition _set_Misa_O (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_O (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4050,18 +3644,14 @@ Definition _set_Misa_O (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_O (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_O (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 14 14 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_N (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 13 13.
+Definition _get_Misa_N (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 13 13.
-Definition _set_Misa_N (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_N (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4071,18 +3661,14 @@ Definition _set_Misa_N (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_N (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_N (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 13 13 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_M (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 12 12.
+Definition _get_Misa_M (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 12 12.
-Definition _set_Misa_M (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_M (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4092,18 +3678,14 @@ Definition _set_Misa_M (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_M (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_M (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 12 12 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_L (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 11 11.
+Definition _get_Misa_L (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 11 11.
-Definition _set_Misa_L (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_L (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4113,18 +3695,14 @@ Definition _set_Misa_L (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_L (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_L (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 11 11 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_K (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 10 10.
+Definition _get_Misa_K (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 10 10.
-Definition _set_Misa_K (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_K (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4134,18 +3712,14 @@ Definition _set_Misa_K (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_K (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_K (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 10 10 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_J (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 9 9.
+Definition _get_Misa_J (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 9 9.
-Definition _set_Misa_J (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_J (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4155,18 +3729,14 @@ Definition _set_Misa_J (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_J (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_J (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 9 9 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_I (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 8 8.
+Definition _get_Misa_I (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 8 8.
-Definition _set_Misa_I (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_I (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4176,18 +3746,14 @@ Definition _set_Misa_I (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_I (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_I (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_H (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 7 7.
+Definition _get_Misa_H (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 7 7.
-Definition _set_Misa_H (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_H (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4197,18 +3763,14 @@ Definition _set_Misa_H (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_H (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_H (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 7 7 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_G (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 6 6.
+Definition _get_Misa_G (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 6 6.
-Definition _set_Misa_G (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_G (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4218,18 +3780,14 @@ Definition _set_Misa_G (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_G (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_G (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 6 6 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_F (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 5 5.
+Definition _get_Misa_F (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 5 5.
-Definition _set_Misa_F (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_F (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4239,18 +3797,14 @@ Definition _set_Misa_F (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_F (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_F (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 5 5 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_E (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 4 4.
+Definition _get_Misa_E (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 4 4.
-Definition _set_Misa_E (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_E (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4260,18 +3814,14 @@ Definition _set_Misa_E (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_E (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_E (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_D (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 3 3.
+Definition _get_Misa_D (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 3 3.
-Definition _set_Misa_D (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_D (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4281,18 +3831,14 @@ Definition _set_Misa_D (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_D (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_D (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 3 3 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_C (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 2 2.
+Definition _get_Misa_C (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 2 2.
-Definition _set_Misa_C (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_C (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4302,18 +3848,14 @@ Definition _set_Misa_C (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_C (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_C (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 2 2 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_B (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 1 1.
+Definition _get_Misa_B (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 1 1.
-Definition _set_Misa_B (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_B (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4323,18 +3865,14 @@ Definition _set_Misa_B (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_B (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_B (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Misa_A (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 0 0.
+Definition _get_Misa_A (v : Misa) : mword 1 := subrange_vec_dec v.(Misa_Misa_chunk_0) 0 0.
-Definition _set_Misa_A (r_ref : register_ref regstate register_value Misa) (v : mword 1)
-: M (unit) :=
-
+Definition _set_Misa_A (r_ref : register_ref regstate register_value Misa) (v : mword 1) : M (unit) :=
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4344,84 +3882,60 @@ Definition _set_Misa_A (r_ref : register_ref regstate register_value Misa) (v :
write_reg r_ref r
: M (unit).
-Definition _update_Misa_A (v : Misa) (x : mword 1)
-: Misa :=
-
+Definition _update_Misa_A (v : Misa) (x : mword 1) : Misa :=
{[ v with
Misa_Misa_chunk_0 :=
(update_subrange_vec_dec v.(Misa_Misa_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition legalize_misa (m : Misa) (v : mword 64)
-: M (Misa) :=
-
- (if ((sys_enable_writable_misa tt)) then
+Definition ext_veto_disable_C '(tt : unit) : M (bool) := returnm false.
+
+Definition legalize_misa (m : Misa) (v : mword 64) : M (Misa) :=
+ (if sys_enable_writable_misa tt then
let v := Mk_Misa v in
- (and_boolM (returnm ((eq_vec (_get_Misa_C v) ((bool_to_bits false) : mword 1)) : bool))
- (((read_reg nextPC_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- (bit_to_bool (access_vec_dec w__0 1)) >>= fun w__1 : bool =>
- returnm ((Bool.eqb w__1 true)
- : bool))) >>= fun w__2 : bool =>
- returnm ((if sumbool_of_bool (w__2) then m
- else _update_Misa_C m (_get_Misa_C v))
- : Misa)
- else returnm (m : Misa))
+ (or_boolM
+ ((and_boolM (returnm ((eq_vec (_get_Misa_C v) ('b"0" : mword 1)) : bool))
+ ((or_boolM
+ (((read_reg nextPC_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ returnm ((eq_bit (access_vec_dec w__0 1) B1) : bool))
+ ((ext_veto_disable_C tt)
+ : M (bool)))
+ : M (bool)))
+ : M (bool)) (returnm ((negb (sys_enable_rvc tt)) : bool))) >>= fun w__4 : bool =>
+ let m := if sumbool_of_bool w__4 then m else _update_Misa_C m (_get_Misa_C v) in
+ returnm (if orb (negb (sys_enable_fdext tt))
+ (andb (eq_vec (_get_Misa_D v) ('b"1" : mword 1))
+ (eq_vec (_get_Misa_F v) ('b"0" : mword 1))) then
+ m
+ else _update_Misa_D (_update_Misa_F m (_get_Misa_F v)) (_get_Misa_D v))
+ else returnm m)
: M (Misa).
-Definition haveAtomics '(tt : unit)
-: M (bool) :=
-
- read_reg misa_ref >>= fun w__0 : Misa =>
- returnm ((eq_vec (_get_Misa_A w__0) ((bool_to_bits true) : mword 1))
- : bool).
+Definition haveAtomics '(tt : unit) : M (bool) :=
+ read_reg misa_ref >>= fun w__0 : Misa => returnm (eq_vec (_get_Misa_A w__0) ('b"1" : mword 1)).
-Definition haveRVC '(tt : unit)
-: M (bool) :=
-
- read_reg misa_ref >>= fun w__0 : Misa =>
- returnm ((eq_vec (_get_Misa_C w__0) ((bool_to_bits true) : mword 1))
- : bool).
+Definition haveRVC '(tt : unit) : M (bool) :=
+ read_reg misa_ref >>= fun w__0 : Misa => returnm (eq_vec (_get_Misa_C w__0) ('b"1" : mword 1)).
-Definition haveMulDiv '(tt : unit)
-: M (bool) :=
-
- read_reg misa_ref >>= fun w__0 : Misa =>
- returnm ((eq_vec (_get_Misa_M w__0) ((bool_to_bits true) : mword 1))
- : bool).
+Definition haveMulDiv '(tt : unit) : M (bool) :=
+ read_reg misa_ref >>= fun w__0 : Misa => returnm (eq_vec (_get_Misa_M w__0) ('b"1" : mword 1)).
-Definition haveSupMode '(tt : unit)
-: M (bool) :=
-
- read_reg misa_ref >>= fun w__0 : Misa =>
- returnm ((eq_vec (_get_Misa_S w__0) ((bool_to_bits true) : mword 1))
- : bool).
+Definition haveSupMode '(tt : unit) : M (bool) :=
+ read_reg misa_ref >>= fun w__0 : Misa => returnm (eq_vec (_get_Misa_S w__0) ('b"1" : mword 1)).
-Definition haveUsrMode '(tt : unit)
-: M (bool) :=
-
- read_reg misa_ref >>= fun w__0 : Misa =>
- returnm ((eq_vec (_get_Misa_U w__0) ((bool_to_bits true) : mword 1))
- : bool).
+Definition haveUsrMode '(tt : unit) : M (bool) :=
+ read_reg misa_ref >>= fun w__0 : Misa => returnm (eq_vec (_get_Misa_U w__0) ('b"1" : mword 1)).
-Definition haveNExt '(tt : unit)
-: M (bool) :=
-
- read_reg misa_ref >>= fun w__0 : Misa =>
- returnm ((eq_vec (_get_Misa_N w__0) ((bool_to_bits true) : mword 1))
- : bool).
+Definition haveNExt '(tt : unit) : M (bool) :=
+ read_reg misa_ref >>= fun w__0 : Misa => returnm (eq_vec (_get_Misa_N w__0) ('b"1" : mword 1)).
-Definition Mk_Mstatus (v : mword 64)
-: Mstatus :=
-
+Definition Mk_Mstatus (v : mword 64) : Mstatus :=
{| Mstatus_Mstatus_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Mstatus_bits (v : Mstatus)
-: mword 64 :=
-
+Definition _get_Mstatus_bits (v : Mstatus) : mword 64 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 63 0.
-Definition _set_Mstatus_bits (r_ref : register_ref regstate register_value Mstatus) (v : mword 64)
+Definition _set_Mstatus_bits (r_ref : register_ref regstate register_value Mstatus) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4431,21 +3945,16 @@ Definition _set_Mstatus_bits (r_ref : register_ref regstate register_value Mstat
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_bits (v : Mstatus) (x : mword 64)
-: Mstatus :=
-
+Definition _update_Mstatus_bits (v : Mstatus) (x : mword 64) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Mstatus_SD (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_SD (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 63 63.
-Definition _set_Mstatus_SD (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_SD (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4455,21 +3964,16 @@ Definition _set_Mstatus_SD (r_ref : register_ref regstate register_value Mstatus
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_SD (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_SD (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 63 63 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_TSR (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_TSR (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 22 22.
-Definition _set_Mstatus_TSR (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_TSR (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4479,21 +3983,16 @@ Definition _set_Mstatus_TSR (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_TSR (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_TSR (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 22 22 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_TW (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_TW (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 21 21.
-Definition _set_Mstatus_TW (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_TW (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4503,21 +4002,16 @@ Definition _set_Mstatus_TW (r_ref : register_ref regstate register_value Mstatus
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_TW (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_TW (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 21 21 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_TVM (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_TVM (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 20 20.
-Definition _set_Mstatus_TVM (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_TVM (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4527,21 +4021,16 @@ Definition _set_Mstatus_TVM (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_TVM (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_TVM (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 20 20 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_MXR (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_MXR (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 19 19.
-Definition _set_Mstatus_MXR (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_MXR (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4551,21 +4040,16 @@ Definition _set_Mstatus_MXR (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_MXR (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_MXR (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 19 19 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_SUM (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_SUM (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 18 18.
-Definition _set_Mstatus_SUM (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_SUM (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4575,21 +4059,16 @@ Definition _set_Mstatus_SUM (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_SUM (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_SUM (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 18 18 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_MPRV (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_MPRV (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 17 17.
-Definition _set_Mstatus_MPRV (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_MPRV (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4599,21 +4078,16 @@ Definition _set_Mstatus_MPRV (r_ref : register_ref regstate register_value Mstat
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_MPRV (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_MPRV (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 17 17 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_XS (v : Mstatus)
-: mword 2 :=
-
+Definition _get_Mstatus_XS (v : Mstatus) : mword 2 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 16 15.
-Definition _set_Mstatus_XS (r_ref : register_ref regstate register_value Mstatus) (v : mword 2)
+Definition _set_Mstatus_XS (r_ref : register_ref regstate register_value Mstatus) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4623,21 +4097,16 @@ Definition _set_Mstatus_XS (r_ref : register_ref regstate register_value Mstatus
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_XS (v : Mstatus) (x : mword 2)
-: Mstatus :=
-
+Definition _update_Mstatus_XS (v : Mstatus) (x : mword 2) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 16 15 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_Mstatus_FS (v : Mstatus)
-: mword 2 :=
-
+Definition _get_Mstatus_FS (v : Mstatus) : mword 2 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 14 13.
-Definition _set_Mstatus_FS (r_ref : register_ref regstate register_value Mstatus) (v : mword 2)
+Definition _set_Mstatus_FS (r_ref : register_ref regstate register_value Mstatus) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4647,21 +4116,16 @@ Definition _set_Mstatus_FS (r_ref : register_ref regstate register_value Mstatus
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_FS (v : Mstatus) (x : mword 2)
-: Mstatus :=
-
+Definition _update_Mstatus_FS (v : Mstatus) (x : mword 2) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 14 13 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_Mstatus_MPP (v : Mstatus)
-: mword 2 :=
-
+Definition _get_Mstatus_MPP (v : Mstatus) : mword 2 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 12 11.
-Definition _set_Mstatus_MPP (r_ref : register_ref regstate register_value Mstatus) (v : mword 2)
+Definition _set_Mstatus_MPP (r_ref : register_ref regstate register_value Mstatus) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4671,21 +4135,16 @@ Definition _set_Mstatus_MPP (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_MPP (v : Mstatus) (x : mword 2)
-: Mstatus :=
-
+Definition _update_Mstatus_MPP (v : Mstatus) (x : mword 2) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 12 11 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_Mstatus_SPP (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_SPP (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 8 8.
-Definition _set_Mstatus_SPP (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_SPP (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4695,21 +4154,16 @@ Definition _set_Mstatus_SPP (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_SPP (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_SPP (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_MPIE (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_MPIE (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 7 7.
-Definition _set_Mstatus_MPIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_MPIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4719,21 +4173,16 @@ Definition _set_Mstatus_MPIE (r_ref : register_ref regstate register_value Mstat
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_MPIE (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_MPIE (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 7 7 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_SPIE (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_SPIE (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 5 5.
-Definition _set_Mstatus_SPIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_SPIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4743,21 +4192,16 @@ Definition _set_Mstatus_SPIE (r_ref : register_ref regstate register_value Mstat
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_SPIE (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_SPIE (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 5 5 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_UPIE (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_UPIE (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 4 4.
-Definition _set_Mstatus_UPIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_UPIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4767,21 +4211,16 @@ Definition _set_Mstatus_UPIE (r_ref : register_ref regstate register_value Mstat
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_UPIE (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_UPIE (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_MIE (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_MIE (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 3 3.
-Definition _set_Mstatus_MIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_MIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4791,21 +4230,16 @@ Definition _set_Mstatus_MIE (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_MIE (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_MIE (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 3 3 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_SIE (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_SIE (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 1 1.
-Definition _set_Mstatus_SIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_SIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4815,21 +4249,16 @@ Definition _set_Mstatus_SIE (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_SIE (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_SIE (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mstatus_UIE (v : Mstatus)
-: mword 1 :=
-
+Definition _get_Mstatus_UIE (v : Mstatus) : mword 1 :=
subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 0 0.
-Definition _set_Mstatus_UIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
+Definition _set_Mstatus_UIE (r_ref : register_ref regstate register_value Mstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4839,101 +4268,92 @@ Definition _set_Mstatus_UIE (r_ref : register_ref regstate register_value Mstatu
write_reg r_ref r
: M (unit).
-Definition _update_Mstatus_UIE (v : Mstatus) (x : mword 1)
-: Mstatus :=
-
+Definition _update_Mstatus_UIE (v : Mstatus) (x : mword 1) : Mstatus :=
{[ v with
Mstatus_Mstatus_chunk_0 :=
(update_subrange_vec_dec v.(Mstatus_Mstatus_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition effectivePrivilege (m : Mstatus) (priv : Privilege)
-: M (Privilege) :=
-
- (if ((eq_vec (_get_Mstatus_MPRV m) ((bool_to_bits true) : mword 1))) then
+Definition effectivePrivilege (t : AccessType unit) (m : Mstatus) (priv : Privilege) : M (Privilege) :=
+ (if andb (generic_neq t (Execute tt)) (eq_vec (_get_Mstatus_MPRV m) ('b"1" : mword 1)) then
read_reg mstatus_ref >>= fun w__0 : Mstatus =>
(privLevel_of_bits (_get_Mstatus_MPP w__0))
: M (Privilege)
else read_reg cur_privilege_ref : M (Privilege))
: M (Privilege).
-Definition get_mstatus_SXL (m : Mstatus)
-: mword 2 :=
-
- subrange_vec_dec (_get_Mstatus_bits m) 35 34.
+Definition get_mstatus_SXL (m : Mstatus) : mword 2 := subrange_vec_dec (_get_Mstatus_bits m) 35 34.
-Definition set_mstatus_SXL (m : Mstatus) (a : mword 2)
-: Mstatus :=
-
+Definition set_mstatus_SXL (m : Mstatus) (a : mword 2) : Mstatus :=
let m := update_subrange_vec_dec (_get_Mstatus_bits m) 35 34 a in
Mk_Mstatus m.
-Definition get_mstatus_UXL (m : Mstatus)
-: mword 2 :=
-
- subrange_vec_dec (_get_Mstatus_bits m) 33 32.
+Definition get_mstatus_UXL (m : Mstatus) : mword 2 := subrange_vec_dec (_get_Mstatus_bits m) 33 32.
-Definition set_mstatus_UXL (m : Mstatus) (a : mword 2)
-: Mstatus :=
-
+Definition set_mstatus_UXL (m : Mstatus) (a : mword 2) : Mstatus :=
let m := update_subrange_vec_dec (_get_Mstatus_bits m) 33 32 a in
Mk_Mstatus m.
-Definition legalize_mstatus (o : Mstatus) (v : mword 64)
-: M (Mstatus) :=
-
+Definition legalize_mstatus (o : Mstatus) (v : mword 64) : M (Mstatus) :=
let m : Mstatus := Mk_Mstatus v in
let m := _update_Mstatus_XS m (extStatus_to_bits Off) in
(or_boolM
((extStatus_of_bits (_get_Mstatus_FS m)) >>= fun w__0 : ExtStatus =>
- returnm ((eq_vec (extStatus_to_bits w__0) ((extStatus_to_bits Dirty) : mword 2))
- : bool))
+ returnm ((generic_eq w__0 Dirty) : bool))
((extStatus_of_bits (_get_Mstatus_XS m)) >>= fun w__1 : ExtStatus =>
- returnm ((eq_vec (extStatus_to_bits w__1) ((extStatus_to_bits Dirty) : mword 2))
- : bool))) >>= fun w__2 : bool =>
- let m := _update_Mstatus_SD m ((bool_to_bits w__2) : mword 1) in
+ returnm ((generic_eq w__1 Dirty) : bool))) >>= fun dirty =>
+ let m := _update_Mstatus_SD m (bool_to_bits dirty) in
let m := set_mstatus_SXL m (get_mstatus_SXL o) in
let m := set_mstatus_UXL m (get_mstatus_UXL o) in
- let m := _update_Mstatus_UPIE m ((bool_to_bits false) : mword 1) in
- let m := _update_Mstatus_UIE m ((bool_to_bits false) : mword 1) in
- returnm (m
- : Mstatus).
-
-Definition cur_Architecture '(tt : unit)
-: M (Architecture) :=
-
+ (haveNExt tt) >>= fun w__2 : bool =>
+ let m :=
+ if sumbool_of_bool (negb w__2) then
+ let m := _update_Mstatus_UPIE m ('b"0" : mword 1) in
+ _update_Mstatus_UIE m ('b"0" : mword 1)
+ else m in
+ (haveUsrMode tt) >>= fun w__3 : bool =>
+ returnm (if sumbool_of_bool (negb w__3) then _update_Mstatus_MPRV m ('b"0" : mword 1) else m).
+
+Definition cur_Architecture '(tt : unit) : M (Architecture) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
(match w__0 with
- | Machine => read_reg misa_ref >>= fun w__1 : Misa => returnm ((_get_Misa_MXL w__1) : mword 2)
- | Supervisor =>
- read_reg mstatus_ref >>= fun w__2 : Mstatus => returnm ((get_mstatus_SXL w__2) : mword 2)
- | User =>
- read_reg mstatus_ref >>= fun w__3 : Mstatus => returnm ((get_mstatus_UXL w__3) : mword 2)
+ | Machine => read_reg misa_ref >>= fun w__1 : Misa => returnm (_get_Misa_MXL w__1)
+ | Supervisor => read_reg mstatus_ref >>= fun w__2 : Mstatus => returnm (get_mstatus_SXL w__2)
+ | User => read_reg mstatus_ref >>= fun w__3 : Mstatus => returnm (get_mstatus_UXL w__3)
end) >>= fun a : arch_xlen =>
(match (architecture a) with
- | Some (a) => returnm (a : Architecture)
+ | Some a => returnm a
| None => (internal_error "Invalid current architecture") : M (Architecture)
end)
: M (Architecture).
-Definition in32BitMode '(tt : unit)
-: M (bool) :=
-
- (cur_Architecture tt) >>= fun w__0 : Architecture => returnm ((generic_eq w__0 RV32) : bool).
+Definition in32BitMode '(tt : unit) : M (bool) :=
+ (cur_Architecture tt) >>= fun w__0 : Architecture => returnm (generic_eq w__0 RV32).
+
+Definition haveFExt '(tt : unit) : M (bool) :=
+ (and_boolM
+ (read_reg misa_ref >>= fun w__0 : Misa =>
+ returnm ((eq_vec (_get_Misa_F w__0) ('b"1" : mword 1)) : bool))
+ (read_reg mstatus_ref >>= fun w__1 : Mstatus =>
+ returnm ((neq_vec (_get_Mstatus_FS w__1) ('b"00" : mword 2)) : bool)))
+ : M (bool).
+
+Definition haveDExt '(tt : unit) : M (bool) :=
+ (and_boolM
+ (read_reg misa_ref >>= fun w__0 : Misa =>
+ returnm ((eq_vec (_get_Misa_D w__0) ('b"1" : mword 1)) : bool))
+ (read_reg mstatus_ref >>= fun w__1 : Mstatus =>
+ returnm ((neq_vec (_get_Mstatus_FS w__1) ('b"00" : mword 2)) : bool)))
+ : M (bool).
-Definition Mk_Minterrupts (v : mword 64)
-: Minterrupts :=
-
+Definition Mk_Minterrupts (v : mword 64) : Minterrupts :=
{| Minterrupts_Minterrupts_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Minterrupts_bits (v : Minterrupts)
-: mword 64 :=
-
+Definition _get_Minterrupts_bits (v : Minterrupts) : mword 64 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 63 0.
Definition _set_Minterrupts_bits
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 64)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4943,22 +4363,17 @@ Definition _set_Minterrupts_bits
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_bits (v : Minterrupts) (x : mword 64)
-: Minterrupts :=
-
+Definition _update_Minterrupts_bits (v : Minterrupts) (x : mword 64) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Minterrupts_MEI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_MEI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 11 11.
Definition _set_Minterrupts_MEI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4968,22 +4383,17 @@ Definition _set_Minterrupts_MEI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_MEI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_MEI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 11 11 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Minterrupts_SEI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_SEI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 9 9.
Definition _set_Minterrupts_SEI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -4993,22 +4403,17 @@ Definition _set_Minterrupts_SEI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_SEI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_SEI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 9 9 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Minterrupts_UEI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_UEI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 8 8.
Definition _set_Minterrupts_UEI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5018,22 +4423,17 @@ Definition _set_Minterrupts_UEI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_UEI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_UEI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Minterrupts_MTI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_MTI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 7 7.
Definition _set_Minterrupts_MTI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5043,22 +4443,17 @@ Definition _set_Minterrupts_MTI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_MTI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_MTI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 7 7 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Minterrupts_STI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_STI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 5 5.
Definition _set_Minterrupts_STI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5068,22 +4463,17 @@ Definition _set_Minterrupts_STI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_STI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_STI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 5 5 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Minterrupts_UTI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_UTI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 4 4.
Definition _set_Minterrupts_UTI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5093,22 +4483,17 @@ Definition _set_Minterrupts_UTI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_UTI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_UTI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Minterrupts_MSI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_MSI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 3 3.
Definition _set_Minterrupts_MSI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5118,22 +4503,17 @@ Definition _set_Minterrupts_MSI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_MSI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_MSI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 3 3 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Minterrupts_SSI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_SSI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 1 1.
Definition _set_Minterrupts_SSI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5143,22 +4523,17 @@ Definition _set_Minterrupts_SSI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_SSI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_SSI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Minterrupts_USI (v : Minterrupts)
-: mword 1 :=
-
+Definition _get_Minterrupts_USI (v : Minterrupts) : mword 1 :=
subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 0 0.
Definition _set_Minterrupts_USI
-(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Minterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5168,31 +4543,24 @@ Definition _set_Minterrupts_USI
write_reg r_ref r
: M (unit).
-Definition _update_Minterrupts_USI (v : Minterrupts) (x : mword 1)
-: Minterrupts :=
-
+Definition _update_Minterrupts_USI (v : Minterrupts) (x : mword 1) : Minterrupts :=
{[ v with
Minterrupts_Minterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Minterrupts_Minterrupts_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition legalize_mip (o : Minterrupts) (v : mword 64)
-: M (Minterrupts) :=
-
+Definition legalize_mip (o : Minterrupts) (v : mword 64) : M (Minterrupts) :=
let v := Mk_Minterrupts v in
let m := _update_Minterrupts_SEI o (_get_Minterrupts_SEI v) in
let m := _update_Minterrupts_STI m (_get_Minterrupts_STI v) in
let m := _update_Minterrupts_SSI m (_get_Minterrupts_SSI v) in
- (haveUsrMode tt) >>= fun w__0 : bool =>
- returnm ((if sumbool_of_bool (w__0) then
- let m := _update_Minterrupts_UEI m (_get_Minterrupts_UEI v) in
- let m := _update_Minterrupts_UTI m (_get_Minterrupts_UTI v) in
- _update_Minterrupts_USI m (_get_Minterrupts_USI v)
- else m)
- : Minterrupts).
-
-Definition legalize_mie (o : Minterrupts) (v : mword 64)
-: M (Minterrupts) :=
-
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool))) >>= fun w__2 : bool =>
+ returnm (if sumbool_of_bool w__2 then
+ let m := _update_Minterrupts_UEI m (_get_Minterrupts_UEI v) in
+ let m := _update_Minterrupts_UTI m (_get_Minterrupts_UTI v) in
+ _update_Minterrupts_USI m (_get_Minterrupts_USI v)
+ else m).
+
+Definition legalize_mie (o : Minterrupts) (v : mword 64) : M (Minterrupts) :=
let v := Mk_Minterrupts v in
let m := _update_Minterrupts_MEI o (_get_Minterrupts_MEI v) in
let m := _update_Minterrupts_MTI m (_get_Minterrupts_MTI v) in
@@ -5200,35 +4568,27 @@ Definition legalize_mie (o : Minterrupts) (v : mword 64)
let m := _update_Minterrupts_SEI m (_get_Minterrupts_SEI v) in
let m := _update_Minterrupts_STI m (_get_Minterrupts_STI v) in
let m := _update_Minterrupts_SSI m (_get_Minterrupts_SSI v) in
- (haveUsrMode tt) >>= fun w__0 : bool =>
- returnm ((if sumbool_of_bool (w__0) then
- let m := _update_Minterrupts_UEI m (_get_Minterrupts_UEI v) in
- let m := _update_Minterrupts_UTI m (_get_Minterrupts_UTI v) in
- _update_Minterrupts_USI m (_get_Minterrupts_USI v)
- else m)
- : Minterrupts).
-
-Definition legalize_mideleg (o : Minterrupts) (v : mword 64)
-: Minterrupts :=
-
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool))) >>= fun w__2 : bool =>
+ returnm (if sumbool_of_bool w__2 then
+ let m := _update_Minterrupts_UEI m (_get_Minterrupts_UEI v) in
+ let m := _update_Minterrupts_UTI m (_get_Minterrupts_UTI v) in
+ _update_Minterrupts_USI m (_get_Minterrupts_USI v)
+ else m).
+
+Definition legalize_mideleg (o : Minterrupts) (v : mword 64) : Minterrupts :=
let m := Mk_Minterrupts v in
- let m := _update_Minterrupts_MEI m ((bool_to_bits false) : mword 1) in
- let m := _update_Minterrupts_MTI m ((bool_to_bits false) : mword 1) in
- _update_Minterrupts_MSI m ((bool_to_bits false) : mword 1).
+ let m := _update_Minterrupts_MEI m ('b"0" : mword 1) in
+ let m := _update_Minterrupts_MTI m ('b"0" : mword 1) in
+ _update_Minterrupts_MSI m ('b"0" : mword 1).
-Definition Mk_Medeleg (v : mword 64)
-: Medeleg :=
-
+Definition Mk_Medeleg (v : mword 64) : Medeleg :=
{| Medeleg_Medeleg_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Medeleg_bits (v : Medeleg)
-: mword 64 :=
-
+Definition _get_Medeleg_bits (v : Medeleg) : mword 64 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 63 0.
-Definition _set_Medeleg_bits (r_ref : register_ref regstate register_value Medeleg) (v : mword 64)
+Definition _set_Medeleg_bits (r_ref : register_ref regstate register_value Medeleg) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5238,22 +4598,17 @@ Definition _set_Medeleg_bits (r_ref : register_ref regstate register_value Medel
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_bits (v : Medeleg) (x : mword 64)
-: Medeleg :=
-
+Definition _update_Medeleg_bits (v : Medeleg) (x : mword 64) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Medeleg_SAMO_Page_Fault (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_SAMO_Page_Fault (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 15 15.
Definition _set_Medeleg_SAMO_Page_Fault
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5263,22 +4618,17 @@ Definition _set_Medeleg_SAMO_Page_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_SAMO_Page_Fault (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_SAMO_Page_Fault (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 15 15 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_Load_Page_Fault (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_Load_Page_Fault (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 13 13.
Definition _set_Medeleg_Load_Page_Fault
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5288,22 +4638,17 @@ Definition _set_Medeleg_Load_Page_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_Load_Page_Fault (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_Load_Page_Fault (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 13 13 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_Fetch_Page_Fault (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_Fetch_Page_Fault (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 12 12.
Definition _set_Medeleg_Fetch_Page_Fault
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5313,22 +4658,17 @@ Definition _set_Medeleg_Fetch_Page_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_Fetch_Page_Fault (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_Fetch_Page_Fault (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 12 12 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_MEnvCall (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_MEnvCall (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 10 10.
Definition _set_Medeleg_MEnvCall
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5338,22 +4678,17 @@ Definition _set_Medeleg_MEnvCall
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_MEnvCall (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_MEnvCall (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 10 10 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_SEnvCall (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_SEnvCall (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 9 9.
Definition _set_Medeleg_SEnvCall
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5363,22 +4698,17 @@ Definition _set_Medeleg_SEnvCall
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_SEnvCall (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_SEnvCall (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 9 9 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_UEnvCall (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_UEnvCall (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 8 8.
Definition _set_Medeleg_UEnvCall
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5388,22 +4718,17 @@ Definition _set_Medeleg_UEnvCall
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_UEnvCall (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_UEnvCall (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_SAMO_Access_Fault (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_SAMO_Access_Fault (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 7 7.
Definition _set_Medeleg_SAMO_Access_Fault
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5413,22 +4738,17 @@ Definition _set_Medeleg_SAMO_Access_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_SAMO_Access_Fault (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_SAMO_Access_Fault (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 7 7 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_SAMO_Addr_Align (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_SAMO_Addr_Align (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 6 6.
Definition _set_Medeleg_SAMO_Addr_Align
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5438,22 +4758,17 @@ Definition _set_Medeleg_SAMO_Addr_Align
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_SAMO_Addr_Align (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_SAMO_Addr_Align (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 6 6 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_Load_Access_Fault (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_Load_Access_Fault (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 5 5.
Definition _set_Medeleg_Load_Access_Fault
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5463,22 +4778,17 @@ Definition _set_Medeleg_Load_Access_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_Load_Access_Fault (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_Load_Access_Fault (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 5 5 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_Load_Addr_Align (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_Load_Addr_Align (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 4 4.
Definition _set_Medeleg_Load_Addr_Align
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5488,22 +4798,17 @@ Definition _set_Medeleg_Load_Addr_Align
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_Load_Addr_Align (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_Load_Addr_Align (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_Breakpoint (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_Breakpoint (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 3 3.
Definition _set_Medeleg_Breakpoint
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5513,22 +4818,17 @@ Definition _set_Medeleg_Breakpoint
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_Breakpoint (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_Breakpoint (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 3 3 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_Illegal_Instr (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_Illegal_Instr (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 2 2.
Definition _set_Medeleg_Illegal_Instr
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5538,22 +4838,17 @@ Definition _set_Medeleg_Illegal_Instr
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_Illegal_Instr (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_Illegal_Instr (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 2 2 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_Fetch_Access_Fault (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_Fetch_Access_Fault (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 1 1.
Definition _set_Medeleg_Fetch_Access_Fault
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5563,22 +4858,17 @@ Definition _set_Medeleg_Fetch_Access_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_Fetch_Access_Fault (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_Fetch_Access_Fault (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Medeleg_Fetch_Addr_Align (v : Medeleg)
-: mword 1 :=
-
+Definition _get_Medeleg_Fetch_Addr_Align (v : Medeleg) : mword 1 :=
subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 0 0.
Definition _set_Medeleg_Fetch_Addr_Align
-(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Medeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5588,32 +4878,22 @@ Definition _set_Medeleg_Fetch_Addr_Align
write_reg r_ref r
: M (unit).
-Definition _update_Medeleg_Fetch_Addr_Align (v : Medeleg) (x : mword 1)
-: Medeleg :=
-
+Definition _update_Medeleg_Fetch_Addr_Align (v : Medeleg) (x : mword 1) : Medeleg :=
{[ v with
Medeleg_Medeleg_chunk_0 :=
(update_subrange_vec_dec v.(Medeleg_Medeleg_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition legalize_medeleg (o : Medeleg) (v : mword 64)
-: Medeleg :=
-
+Definition legalize_medeleg (o : Medeleg) (v : mword 64) : Medeleg :=
let m := Mk_Medeleg v in
- _update_Medeleg_MEnvCall m ((bool_to_bits false) : mword 1).
+ _update_Medeleg_MEnvCall m ('b"0" : mword 1).
-Definition Mk_Mtvec (v : mword 64)
-: Mtvec :=
-
+Definition Mk_Mtvec (v : mword 64) : Mtvec :=
{| Mtvec_Mtvec_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Mtvec_bits (v : Mtvec)
-: mword 64 :=
-
- subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 63 0.
+Definition _get_Mtvec_bits (v : Mtvec) : mword 64 := subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 63 0.
-Definition _set_Mtvec_bits (r_ref : register_ref regstate register_value Mtvec) (v : mword 64)
+Definition _set_Mtvec_bits (r_ref : register_ref regstate register_value Mtvec) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5623,21 +4903,15 @@ Definition _set_Mtvec_bits (r_ref : register_ref regstate register_value Mtvec)
write_reg r_ref r
: M (unit).
-Definition _update_Mtvec_bits (v : Mtvec) (x : mword 64)
-: Mtvec :=
-
+Definition _update_Mtvec_bits (v : Mtvec) (x : mword 64) : Mtvec :=
{[ v with
Mtvec_Mtvec_chunk_0 :=
(update_subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Mtvec_Base (v : Mtvec)
-: mword 62 :=
-
- subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 63 2.
+Definition _get_Mtvec_Base (v : Mtvec) : mword 62 := subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 63 2.
-Definition _set_Mtvec_Base (r_ref : register_ref regstate register_value Mtvec) (v : mword 62)
+Definition _set_Mtvec_Base (r_ref : register_ref regstate register_value Mtvec) (v : mword 62)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5647,18 +4921,15 @@ Definition _set_Mtvec_Base (r_ref : register_ref regstate register_value Mtvec)
write_reg r_ref r
: M (unit).
-Definition _update_Mtvec_Base (v : Mtvec) (x : mword 62)
-: Mtvec :=
-
+Definition _update_Mtvec_Base (v : Mtvec) (x : mword 62) : Mtvec :=
{[ v with
Mtvec_Mtvec_chunk_0 :=
(update_subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 63 2 (subrange_vec_dec x 61 0)) ]}.
-Definition _get_Mtvec_Mode (v : Mtvec) : mword 2 := subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 1 0.
+Definition _get_Mtvec_Mode (v : Mtvec) : mword 2 := subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 1 0.
-Definition _set_Mtvec_Mode (r_ref : register_ref regstate register_value Mtvec) (v : mword 2)
+Definition _set_Mtvec_Mode (r_ref : register_ref regstate register_value Mtvec) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5668,16 +4939,12 @@ Definition _set_Mtvec_Mode (r_ref : register_ref regstate register_value Mtvec)
write_reg r_ref r
: M (unit).
-Definition _update_Mtvec_Mode (v : Mtvec) (x : mword 2)
-: Mtvec :=
-
+Definition _update_Mtvec_Mode (v : Mtvec) (x : mword 2) : Mtvec :=
{[ v with
Mtvec_Mtvec_chunk_0 :=
(update_subrange_vec_dec v.(Mtvec_Mtvec_chunk_0) 1 0 (subrange_vec_dec x 1 0)) ]}.
-Definition legalize_tvec (o : Mtvec) (v : mword 64)
-: Mtvec :=
-
+Definition legalize_tvec (o : Mtvec) (v : mword 64) : Mtvec :=
let v := Mk_Mtvec v in
match (trapVectorMode_of_bits (_get_Mtvec_Mode v)) with
| TV_Direct => v
@@ -5685,19 +4952,14 @@ Definition legalize_tvec (o : Mtvec) (v : mword 64)
| _ => _update_Mtvec_Mode v (_get_Mtvec_Mode o)
end.
-Definition Mk_Mcause (v : mword 64)
-: Mcause :=
-
+Definition Mk_Mcause (v : mword 64) : Mcause :=
{| Mcause_Mcause_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Mcause_bits (v : Mcause)
-: mword 64 :=
-
+Definition _get_Mcause_bits (v : Mcause) : mword 64 :=
subrange_vec_dec v.(Mcause_Mcause_chunk_0) 63 0.
-Definition _set_Mcause_bits (r_ref : register_ref regstate register_value Mcause) (v : mword 64)
+Definition _set_Mcause_bits (r_ref : register_ref regstate register_value Mcause) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5707,22 +4969,17 @@ Definition _set_Mcause_bits (r_ref : register_ref regstate register_value Mcause
write_reg r_ref r
: M (unit).
-Definition _update_Mcause_bits (v : Mcause) (x : mword 64)
-: Mcause :=
-
+Definition _update_Mcause_bits (v : Mcause) (x : mword 64) : Mcause :=
{[ v with
Mcause_Mcause_chunk_0 :=
(update_subrange_vec_dec v.(Mcause_Mcause_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Mcause_IsInterrupt (v : Mcause)
-: mword 1 :=
-
+Definition _get_Mcause_IsInterrupt (v : Mcause) : mword 1 :=
subrange_vec_dec v.(Mcause_Mcause_chunk_0) 63 63.
Definition _set_Mcause_IsInterrupt
-(r_ref : register_ref regstate register_value Mcause) (v : mword 1)
+(r_ref : register_ref regstate register_value Mcause) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5732,21 +4989,16 @@ Definition _set_Mcause_IsInterrupt
write_reg r_ref r
: M (unit).
-Definition _update_Mcause_IsInterrupt (v : Mcause) (x : mword 1)
-: Mcause :=
-
+Definition _update_Mcause_IsInterrupt (v : Mcause) (x : mword 1) : Mcause :=
{[ v with
Mcause_Mcause_chunk_0 :=
(update_subrange_vec_dec v.(Mcause_Mcause_chunk_0) 63 63 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Mcause_Cause (v : Mcause)
-: mword 63 :=
-
+Definition _get_Mcause_Cause (v : Mcause) : mword 63 :=
subrange_vec_dec v.(Mcause_Mcause_chunk_0) 62 0.
-Definition _set_Mcause_Cause (r_ref : register_ref regstate register_value Mcause) (v : mword 63)
+Definition _set_Mcause_Cause (r_ref : register_ref regstate register_value Mcause) (v : mword 63)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5756,63 +5008,45 @@ Definition _set_Mcause_Cause (r_ref : register_ref regstate register_value Mcaus
write_reg r_ref r
: M (unit).
-Definition _update_Mcause_Cause (v : Mcause) (x : mword 63)
-: Mcause :=
-
+Definition _update_Mcause_Cause (v : Mcause) (x : mword 63) : Mcause :=
{[ v with
Mcause_Mcause_chunk_0 :=
(update_subrange_vec_dec v.(Mcause_Mcause_chunk_0) 62 0 (subrange_vec_dec x 62 0)) ]}.
-Definition tvec_addr (m : Mtvec) (c : Mcause)
-: option (mword 64) :=
-
- let base : xlenbits := concat_vec (_get_Mtvec_Base m) (vec_of_bits [B0;B0] : mword 2) in
+Definition tvec_addr (m : Mtvec) (c : Mcause) : option (mword 64) :=
+ let base : xlenbits := concat_vec (_get_Mtvec_Base m) ('b"00" : mword 2) in
match (trapVectorMode_of_bits (_get_Mtvec_Mode m)) with
- | TV_Direct => Some (base)
+ | TV_Direct => Some base
| TV_Vector =>
- if ((eq_vec (_get_Mcause_IsInterrupt c) ((bool_to_bits true) : mword 1))) then
- Some
- (add_vec base (shiftl (EXTZ 64 (_get_Mcause_Cause c)) 2))
- else Some (base)
+ if eq_vec (_get_Mcause_IsInterrupt c) ('b"1" : mword 1) then
+ Some (add_vec base (shiftl (EXTZ 64 (_get_Mcause_Cause c)) 2))
+ else Some base
| TV_Reserved => None
end.
-Definition legalize_xepc (v : mword 64)
-: M (mword 64) :=
-
- (or_boolM (returnm ((sys_enable_writable_misa tt) : bool))
+Definition legalize_xepc (v : mword 64) : M (mword 64) :=
+ (or_boolM (returnm ((andb (sys_enable_writable_misa tt) (sys_enable_rvc tt)) : bool))
(read_reg misa_ref >>= fun w__0 : Misa =>
- returnm ((eq_vec (_get_Misa_C w__0) ((bool_to_bits true) : mword 1))
- : bool))) >>= fun w__1 : bool =>
- returnm ((if sumbool_of_bool (w__1) then update_vec_dec v 0 B0
- else and_vec v (EXTS 64 (vec_of_bits [B1;B0;B0] : mword 3)))
- : mword 64).
+ returnm ((eq_vec (_get_Misa_C w__0) ('b"1" : mword 1)) : bool))) >>= fun w__1 : bool =>
+ returnm (if sumbool_of_bool w__1 then update_vec_dec v 0 B0
+ else and_vec v (EXTS 64 ('b"100" : mword 3))).
-Definition pc_alignment_mask '(tt : unit)
-: M (mword 64) :=
-
+Definition pc_alignment_mask '(tt : unit) : M (mword 64) :=
read_reg misa_ref >>= fun w__0 : Misa =>
- returnm ((not_vec
- (EXTZ 64
- (if ((eq_vec (_get_Misa_C w__0) ((bool_to_bits true) : mword 1))) then
- (vec_of_bits [B0;B0] : mword 2)
- else (vec_of_bits [B1;B0] : mword 2))))
- : mword 64).
-
-Definition Mk_Counteren (v : mword 32)
-: Counteren :=
-
+ returnm (not_vec
+ (EXTZ 64
+ (if eq_vec (_get_Misa_C w__0) ('b"1" : mword 1) then 'b"00" : mword 2
+ else 'b"10" : mword 2))).
+
+Definition Mk_Counteren (v : mword 32) : Counteren :=
{| Counteren_Counteren_chunk_0 := (subrange_vec_dec v 31 0) |}.
-Definition _get_Counteren_bits (v : Counteren)
-: mword 32 :=
-
+Definition _get_Counteren_bits (v : Counteren) : mword 32 :=
subrange_vec_dec v.(Counteren_Counteren_chunk_0) 31 0.
Definition _set_Counteren_bits
-(r_ref : register_ref regstate register_value Counteren) (v : mword 32)
+(r_ref : register_ref regstate register_value Counteren) (v : mword 32)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5822,22 +5056,17 @@ Definition _set_Counteren_bits
write_reg r_ref r
: M (unit).
-Definition _update_Counteren_bits (v : Counteren) (x : mword 32)
-: Counteren :=
-
+Definition _update_Counteren_bits (v : Counteren) (x : mword 32) : Counteren :=
{[ v with
Counteren_Counteren_chunk_0 :=
(update_subrange_vec_dec v.(Counteren_Counteren_chunk_0) 31 0 (subrange_vec_dec x 31 0)) ]}.
-Definition _get_Counteren_HPM (v : Counteren)
-: mword 29 :=
-
+Definition _get_Counteren_HPM (v : Counteren) : mword 29 :=
subrange_vec_dec v.(Counteren_Counteren_chunk_0) 31 3.
Definition _set_Counteren_HPM
-(r_ref : register_ref regstate register_value Counteren) (v : mword 29)
+(r_ref : register_ref regstate register_value Counteren) (v : mword 29)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5847,21 +5076,16 @@ Definition _set_Counteren_HPM
write_reg r_ref r
: M (unit).
-Definition _update_Counteren_HPM (v : Counteren) (x : mword 29)
-: Counteren :=
-
+Definition _update_Counteren_HPM (v : Counteren) (x : mword 29) : Counteren :=
{[ v with
Counteren_Counteren_chunk_0 :=
(update_subrange_vec_dec v.(Counteren_Counteren_chunk_0) 31 3 (subrange_vec_dec x 28 0)) ]}.
-Definition _get_Counteren_IR (v : Counteren)
-: mword 1 :=
-
+Definition _get_Counteren_IR (v : Counteren) : mword 1 :=
subrange_vec_dec v.(Counteren_Counteren_chunk_0) 2 2.
-Definition _set_Counteren_IR (r_ref : register_ref regstate register_value Counteren) (v : mword 1)
+Definition _set_Counteren_IR (r_ref : register_ref regstate register_value Counteren) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5871,21 +5095,16 @@ Definition _set_Counteren_IR (r_ref : register_ref regstate register_value Count
write_reg r_ref r
: M (unit).
-Definition _update_Counteren_IR (v : Counteren) (x : mword 1)
-: Counteren :=
-
+Definition _update_Counteren_IR (v : Counteren) (x : mword 1) : Counteren :=
{[ v with
Counteren_Counteren_chunk_0 :=
(update_subrange_vec_dec v.(Counteren_Counteren_chunk_0) 2 2 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Counteren_TM (v : Counteren)
-: mword 1 :=
-
+Definition _get_Counteren_TM (v : Counteren) : mword 1 :=
subrange_vec_dec v.(Counteren_Counteren_chunk_0) 1 1.
-Definition _set_Counteren_TM (r_ref : register_ref regstate register_value Counteren) (v : mword 1)
+Definition _set_Counteren_TM (r_ref : register_ref regstate register_value Counteren) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5895,21 +5114,16 @@ Definition _set_Counteren_TM (r_ref : register_ref regstate register_value Count
write_reg r_ref r
: M (unit).
-Definition _update_Counteren_TM (v : Counteren) (x : mword 1)
-: Counteren :=
-
+Definition _update_Counteren_TM (v : Counteren) (x : mword 1) : Counteren :=
{[ v with
Counteren_Counteren_chunk_0 :=
(update_subrange_vec_dec v.(Counteren_Counteren_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Counteren_CY (v : Counteren)
-: mword 1 :=
-
+Definition _get_Counteren_CY (v : Counteren) : mword 1 :=
subrange_vec_dec v.(Counteren_Counteren_chunk_0) 0 0.
-Definition _set_Counteren_CY (r_ref : register_ref regstate register_value Counteren) (v : mword 1)
+Definition _set_Counteren_CY (r_ref : register_ref regstate register_value Counteren) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5919,61 +5133,103 @@ Definition _set_Counteren_CY (r_ref : register_ref regstate register_value Count
write_reg r_ref r
: M (unit).
-Definition _update_Counteren_CY (v : Counteren) (x : mword 1)
-: Counteren :=
-
+Definition _update_Counteren_CY (v : Counteren) (x : mword 1) : Counteren :=
{[ v with
Counteren_Counteren_chunk_0 :=
(update_subrange_vec_dec v.(Counteren_Counteren_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition legalize_mcounteren (c : Counteren) (v : mword 64)
-: M (Counteren) :=
-
- (cast_unit_vec (access_vec_dec v 2)) >>= fun w__0 : mword 1 =>
- let c := _update_Counteren_IR c (w__0 : mword 1) in
- (cast_unit_vec (access_vec_dec v 1)) >>= fun w__1 : mword 1 =>
- let c := _update_Counteren_TM c (w__1 : mword 1) in
- (cast_unit_vec (access_vec_dec v 0)) >>= fun w__2 : mword 1 =>
- let c := _update_Counteren_CY c (w__2 : mword 1) in
- returnm (c
- : Counteren).
-
-Definition legalize_scounteren (c : Counteren) (v : mword 64)
-: M (Counteren) :=
-
- (cast_unit_vec (access_vec_dec v 2)) >>= fun w__0 : mword 1 =>
- let c := _update_Counteren_IR c (w__0 : mword 1) in
- (cast_unit_vec (access_vec_dec v 1)) >>= fun w__1 : mword 1 =>
- let c := _update_Counteren_TM c (w__1 : mword 1) in
- (cast_unit_vec (access_vec_dec v 0)) >>= fun w__2 : mword 1 =>
- let c := _update_Counteren_CY c (w__2 : mword 1) in
- returnm (c
- : Counteren).
-
-Definition retire_instruction '(tt : unit)
+Definition legalize_mcounteren (c : Counteren) (v : mword 64) : Counteren :=
+ let c := _update_Counteren_IR c (vec_of_bits [access_vec_dec v 2] : mword 1) in
+ let c := _update_Counteren_TM c (vec_of_bits [access_vec_dec v 1] : mword 1) in
+ _update_Counteren_CY c (vec_of_bits [access_vec_dec v 0] : mword 1).
+
+Definition legalize_scounteren (c : Counteren) (v : mword 64) : Counteren :=
+ let c := _update_Counteren_IR c (vec_of_bits [access_vec_dec v 2] : mword 1) in
+ let c := _update_Counteren_TM c (vec_of_bits [access_vec_dec v 1] : mword 1) in
+ _update_Counteren_CY c (vec_of_bits [access_vec_dec v 0] : mword 1).
+
+Definition Mk_Counterin (v : mword 32) : Counterin :=
+ {| Counterin_Counterin_chunk_0 := (subrange_vec_dec v 31 0) |}.
+
+Definition _get_Counterin_bits (v : Counterin) : mword 32 :=
+ subrange_vec_dec v.(Counterin_Counterin_chunk_0) 31 0.
+
+Definition _set_Counterin_bits
+(r_ref : register_ref regstate register_value Counterin) (v : mword 32)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Counterin_Counterin_chunk_0 :=
+ (update_subrange_vec_dec r.(Counterin_Counterin_chunk_0) 31 0 (subrange_vec_dec v 31 0)) ]}
+ : Counterin in
+ write_reg r_ref r
+ : M (unit).
+
+Definition _update_Counterin_bits (v : Counterin) (x : mword 32) : Counterin :=
+ {[ v with
+ Counterin_Counterin_chunk_0 :=
+ (update_subrange_vec_dec v.(Counterin_Counterin_chunk_0) 31 0 (subrange_vec_dec x 31 0)) ]}.
+
+Definition _get_Counterin_IR (v : Counterin) : mword 1 :=
+ subrange_vec_dec v.(Counterin_Counterin_chunk_0) 2 2.
+
+Definition _set_Counterin_IR (r_ref : register_ref regstate register_value Counterin) (v : mword 1)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Counterin_Counterin_chunk_0 :=
+ (update_subrange_vec_dec r.(Counterin_Counterin_chunk_0) 2 2 (subrange_vec_dec v 0 0)) ]}
+ : Counterin in
+ write_reg r_ref r
+ : M (unit).
+
+Definition _update_Counterin_IR (v : Counterin) (x : mword 1) : Counterin :=
+ {[ v with
+ Counterin_Counterin_chunk_0 :=
+ (update_subrange_vec_dec v.(Counterin_Counterin_chunk_0) 2 2 (subrange_vec_dec x 0 0)) ]}.
+
+Definition _get_Counterin_CY (v : Counterin) : mword 1 :=
+ subrange_vec_dec v.(Counterin_Counterin_chunk_0) 0 0.
+
+Definition _set_Counterin_CY (r_ref : register_ref regstate register_value Counterin) (v : mword 1)
: M (unit) :=
-
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Counterin_Counterin_chunk_0 :=
+ (update_subrange_vec_dec r.(Counterin_Counterin_chunk_0) 0 0 (subrange_vec_dec v 0 0)) ]}
+ : Counterin in
+ write_reg r_ref r
+ : M (unit).
+
+Definition _update_Counterin_CY (v : Counterin) (x : mword 1) : Counterin :=
+ {[ v with
+ Counterin_Counterin_chunk_0 :=
+ (update_subrange_vec_dec v.(Counterin_Counterin_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
+
+Definition legalize_mcountinhibit (c : Counterin) (v : mword 64) : Counterin :=
+ let c := _update_Counterin_IR c (vec_of_bits [access_vec_dec v 2] : mword 1) in
+ _update_Counterin_CY c (vec_of_bits [access_vec_dec v 0] : mword 1).
+
+Definition retire_instruction '(tt : unit) : M (unit) :=
read_reg minstret_written_ref >>= fun w__0 : bool =>
- (if ((Bool.eqb w__0 true)) then write_reg minstret_written_ref false : M (unit)
+ (if Bool.eqb w__0 true then write_reg minstret_written_ref false : M (unit)
else
((read_reg minstret_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
write_reg minstret_ref (add_vec_int w__1 1)
: M (unit))
: M (unit).
-Definition Mk_Sstatus (v : mword 64)
-: Sstatus :=
-
+Definition Mk_Sstatus (v : mword 64) : Sstatus :=
{| Sstatus_Sstatus_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Sstatus_bits (v : Sstatus)
-: mword 64 :=
-
+Definition _get_Sstatus_bits (v : Sstatus) : mword 64 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 63 0.
-Definition _set_Sstatus_bits (r_ref : register_ref regstate register_value Sstatus) (v : mword 64)
+Definition _set_Sstatus_bits (r_ref : register_ref regstate register_value Sstatus) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -5983,21 +5239,16 @@ Definition _set_Sstatus_bits (r_ref : register_ref regstate register_value Sstat
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_bits (v : Sstatus) (x : mword 64)
-: Sstatus :=
-
+Definition _update_Sstatus_bits (v : Sstatus) (x : mword 64) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Sstatus_SD (v : Sstatus)
-: mword 1 :=
-
+Definition _get_Sstatus_SD (v : Sstatus) : mword 1 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 63 63.
-Definition _set_Sstatus_SD (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
+Definition _set_Sstatus_SD (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6007,21 +5258,16 @@ Definition _set_Sstatus_SD (r_ref : register_ref regstate register_value Sstatus
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_SD (v : Sstatus) (x : mword 1)
-: Sstatus :=
-
+Definition _update_Sstatus_SD (v : Sstatus) (x : mword 1) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 63 63 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sstatus_MXR (v : Sstatus)
-: mword 1 :=
-
+Definition _get_Sstatus_MXR (v : Sstatus) : mword 1 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 19 19.
-Definition _set_Sstatus_MXR (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
+Definition _set_Sstatus_MXR (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6031,21 +5277,16 @@ Definition _set_Sstatus_MXR (r_ref : register_ref regstate register_value Sstatu
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_MXR (v : Sstatus) (x : mword 1)
-: Sstatus :=
-
+Definition _update_Sstatus_MXR (v : Sstatus) (x : mword 1) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 19 19 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sstatus_SUM (v : Sstatus)
-: mword 1 :=
-
+Definition _get_Sstatus_SUM (v : Sstatus) : mword 1 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 18 18.
-Definition _set_Sstatus_SUM (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
+Definition _set_Sstatus_SUM (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6055,21 +5296,16 @@ Definition _set_Sstatus_SUM (r_ref : register_ref regstate register_value Sstatu
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_SUM (v : Sstatus) (x : mword 1)
-: Sstatus :=
-
+Definition _update_Sstatus_SUM (v : Sstatus) (x : mword 1) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 18 18 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sstatus_XS (v : Sstatus)
-: mword 2 :=
-
+Definition _get_Sstatus_XS (v : Sstatus) : mword 2 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 16 15.
-Definition _set_Sstatus_XS (r_ref : register_ref regstate register_value Sstatus) (v : mword 2)
+Definition _set_Sstatus_XS (r_ref : register_ref regstate register_value Sstatus) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6079,21 +5315,16 @@ Definition _set_Sstatus_XS (r_ref : register_ref regstate register_value Sstatus
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_XS (v : Sstatus) (x : mword 2)
-: Sstatus :=
-
+Definition _update_Sstatus_XS (v : Sstatus) (x : mword 2) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 16 15 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_Sstatus_FS (v : Sstatus)
-: mword 2 :=
-
+Definition _get_Sstatus_FS (v : Sstatus) : mword 2 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 14 13.
-Definition _set_Sstatus_FS (r_ref : register_ref regstate register_value Sstatus) (v : mword 2)
+Definition _set_Sstatus_FS (r_ref : register_ref regstate register_value Sstatus) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6103,21 +5334,16 @@ Definition _set_Sstatus_FS (r_ref : register_ref regstate register_value Sstatus
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_FS (v : Sstatus) (x : mword 2)
-: Sstatus :=
-
+Definition _update_Sstatus_FS (v : Sstatus) (x : mword 2) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 14 13 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_Sstatus_SPP (v : Sstatus)
-: mword 1 :=
-
+Definition _get_Sstatus_SPP (v : Sstatus) : mword 1 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 8 8.
-Definition _set_Sstatus_SPP (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
+Definition _set_Sstatus_SPP (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6127,21 +5353,16 @@ Definition _set_Sstatus_SPP (r_ref : register_ref regstate register_value Sstatu
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_SPP (v : Sstatus) (x : mword 1)
-: Sstatus :=
-
+Definition _update_Sstatus_SPP (v : Sstatus) (x : mword 1) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sstatus_SPIE (v : Sstatus)
-: mword 1 :=
-
+Definition _get_Sstatus_SPIE (v : Sstatus) : mword 1 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 5 5.
-Definition _set_Sstatus_SPIE (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
+Definition _set_Sstatus_SPIE (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6151,21 +5372,16 @@ Definition _set_Sstatus_SPIE (r_ref : register_ref regstate register_value Sstat
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_SPIE (v : Sstatus) (x : mword 1)
-: Sstatus :=
-
+Definition _update_Sstatus_SPIE (v : Sstatus) (x : mword 1) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 5 5 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sstatus_UPIE (v : Sstatus)
-: mword 1 :=
-
+Definition _get_Sstatus_UPIE (v : Sstatus) : mword 1 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 4 4.
-Definition _set_Sstatus_UPIE (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
+Definition _set_Sstatus_UPIE (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6175,21 +5391,16 @@ Definition _set_Sstatus_UPIE (r_ref : register_ref regstate register_value Sstat
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_UPIE (v : Sstatus) (x : mword 1)
-: Sstatus :=
-
+Definition _update_Sstatus_UPIE (v : Sstatus) (x : mword 1) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sstatus_SIE (v : Sstatus)
-: mword 1 :=
-
+Definition _get_Sstatus_SIE (v : Sstatus) : mword 1 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 1 1.
-Definition _set_Sstatus_SIE (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
+Definition _set_Sstatus_SIE (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6199,21 +5410,16 @@ Definition _set_Sstatus_SIE (r_ref : register_ref regstate register_value Sstatu
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_SIE (v : Sstatus) (x : mword 1)
-: Sstatus :=
-
+Definition _update_Sstatus_SIE (v : Sstatus) (x : mword 1) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sstatus_UIE (v : Sstatus)
-: mword 1 :=
-
+Definition _get_Sstatus_UIE (v : Sstatus) : mword 1 :=
subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 0 0.
-Definition _set_Sstatus_UIE (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
+Definition _set_Sstatus_UIE (r_ref : register_ref regstate register_value Sstatus) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6223,30 +5429,22 @@ Definition _set_Sstatus_UIE (r_ref : register_ref regstate register_value Sstatu
write_reg r_ref r
: M (unit).
-Definition _update_Sstatus_UIE (v : Sstatus) (x : mword 1)
-: Sstatus :=
-
+Definition _update_Sstatus_UIE (v : Sstatus) (x : mword 1) : Sstatus :=
{[ v with
Sstatus_Sstatus_chunk_0 :=
(update_subrange_vec_dec v.(Sstatus_Sstatus_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition get_sstatus_UXL (s : Sstatus)
-: mword 2 :=
-
+Definition get_sstatus_UXL (s : Sstatus) : mword 2 :=
let m := Mk_Mstatus (_get_Sstatus_bits s) in
get_mstatus_UXL m.
-Definition set_sstatus_UXL (s : Sstatus) (a : mword 2)
-: Sstatus :=
-
+Definition set_sstatus_UXL (s : Sstatus) (a : mword 2) : Sstatus :=
let m := Mk_Mstatus (_get_Sstatus_bits s) in
let m := set_mstatus_UXL m a in
Mk_Sstatus (_get_Mstatus_bits m).
-Definition lower_mstatus (m : Mstatus)
-: Sstatus :=
-
- let s := Mk_Sstatus (EXTZ 64 (vec_of_bits [B0] : mword 1)) in
+Definition lower_mstatus (m : Mstatus) : Sstatus :=
+ let s := Mk_Sstatus (EXTZ 64 ('b"0" : mword 1)) in
let s := _update_Sstatus_SD s (_get_Mstatus_SD m) in
let s := set_sstatus_UXL s (get_mstatus_UXL m) in
let s := _update_Sstatus_MXR s (_get_Mstatus_MXR m) in
@@ -6259,48 +5457,35 @@ Definition lower_mstatus (m : Mstatus)
let s := _update_Sstatus_SIE s (_get_Mstatus_SIE m) in
_update_Sstatus_UIE s (_get_Mstatus_UIE m).
-Definition lift_sstatus (m : Mstatus) (s : Sstatus)
-: M (Mstatus) :=
-
+Definition lift_sstatus (m : Mstatus) (s : Sstatus) : M (Mstatus) :=
let m := _update_Mstatus_MXR m (_get_Sstatus_MXR s) in
let m := _update_Mstatus_SUM m (_get_Sstatus_SUM s) in
let m := _update_Mstatus_XS m (_get_Sstatus_XS s) in
let m := _update_Mstatus_FS m (_get_Sstatus_FS s) in
(or_boolM
((extStatus_of_bits (_get_Mstatus_FS m)) >>= fun w__0 : ExtStatus =>
- returnm ((eq_vec (extStatus_to_bits w__0) ((extStatus_to_bits Dirty) : mword 2))
- : bool))
+ returnm ((generic_eq w__0 Dirty) : bool))
((extStatus_of_bits (_get_Mstatus_XS m)) >>= fun w__1 : ExtStatus =>
- returnm ((eq_vec (extStatus_to_bits w__1) ((extStatus_to_bits Dirty) : mword 2))
- : bool))) >>= fun w__2 : bool =>
- let m := _update_Mstatus_SD m ((bool_to_bits w__2) : mword 1) in
+ returnm ((generic_eq w__1 Dirty) : bool))) >>= fun dirty =>
+ let m := _update_Mstatus_SD m (bool_to_bits dirty) in
let m := _update_Mstatus_SPP m (_get_Sstatus_SPP s) in
let m := _update_Mstatus_SPIE m (_get_Sstatus_SPIE s) in
let m := _update_Mstatus_UPIE m (_get_Sstatus_UPIE s) in
let m := _update_Mstatus_SIE m (_get_Sstatus_SIE s) in
let m := _update_Mstatus_UIE m (_get_Sstatus_UIE s) in
- returnm (m
- : Mstatus).
-
-Definition legalize_sstatus (m : Mstatus) (v : mword 64)
-: M (Mstatus) :=
-
- (lift_sstatus m (Mk_Sstatus v))
- : M (Mstatus).
-
-Definition Mk_Sedeleg (v : mword 64)
-: Sedeleg :=
-
+ returnm m.
+
+Definition legalize_sstatus (m : Mstatus) (v : mword 64) : M (Mstatus) :=
+ (lift_sstatus m (Mk_Sstatus v)) : M (Mstatus).
+
+Definition Mk_Sedeleg (v : mword 64) : Sedeleg :=
{| Sedeleg_Sedeleg_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Sedeleg_bits (v : Sedeleg)
-: mword 64 :=
-
+Definition _get_Sedeleg_bits (v : Sedeleg) : mword 64 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 63 0.
-Definition _set_Sedeleg_bits (r_ref : register_ref regstate register_value Sedeleg) (v : mword 64)
+Definition _set_Sedeleg_bits (r_ref : register_ref regstate register_value Sedeleg) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6310,22 +5495,17 @@ Definition _set_Sedeleg_bits (r_ref : register_ref regstate register_value Sedel
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_bits (v : Sedeleg) (x : mword 64)
-: Sedeleg :=
-
+Definition _update_Sedeleg_bits (v : Sedeleg) (x : mword 64) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Sedeleg_UEnvCall (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_UEnvCall (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 8 8.
Definition _set_Sedeleg_UEnvCall
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6335,22 +5515,17 @@ Definition _set_Sedeleg_UEnvCall
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_UEnvCall (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_UEnvCall (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sedeleg_SAMO_Access_Fault (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_SAMO_Access_Fault (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 7 7.
Definition _set_Sedeleg_SAMO_Access_Fault
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6360,22 +5535,17 @@ Definition _set_Sedeleg_SAMO_Access_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_SAMO_Access_Fault (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_SAMO_Access_Fault (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 7 7 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sedeleg_SAMO_Addr_Align (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_SAMO_Addr_Align (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 6 6.
Definition _set_Sedeleg_SAMO_Addr_Align
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6385,22 +5555,17 @@ Definition _set_Sedeleg_SAMO_Addr_Align
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_SAMO_Addr_Align (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_SAMO_Addr_Align (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 6 6 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sedeleg_Load_Access_Fault (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_Load_Access_Fault (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 5 5.
Definition _set_Sedeleg_Load_Access_Fault
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6410,22 +5575,17 @@ Definition _set_Sedeleg_Load_Access_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_Load_Access_Fault (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_Load_Access_Fault (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 5 5 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sedeleg_Load_Addr_Align (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_Load_Addr_Align (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 4 4.
Definition _set_Sedeleg_Load_Addr_Align
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6435,22 +5595,17 @@ Definition _set_Sedeleg_Load_Addr_Align
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_Load_Addr_Align (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_Load_Addr_Align (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sedeleg_Breakpoint (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_Breakpoint (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 3 3.
Definition _set_Sedeleg_Breakpoint
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6460,22 +5615,17 @@ Definition _set_Sedeleg_Breakpoint
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_Breakpoint (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_Breakpoint (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 3 3 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sedeleg_Illegal_Instr (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_Illegal_Instr (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 2 2.
Definition _set_Sedeleg_Illegal_Instr
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6485,22 +5635,17 @@ Definition _set_Sedeleg_Illegal_Instr
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_Illegal_Instr (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_Illegal_Instr (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 2 2 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sedeleg_Fetch_Access_Fault (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_Fetch_Access_Fault (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 1 1.
Definition _set_Sedeleg_Fetch_Access_Fault
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6510,22 +5655,17 @@ Definition _set_Sedeleg_Fetch_Access_Fault
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_Fetch_Access_Fault (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_Fetch_Access_Fault (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sedeleg_Fetch_Addr_Align (v : Sedeleg)
-: mword 1 :=
-
+Definition _get_Sedeleg_Fetch_Addr_Align (v : Sedeleg) : mword 1 :=
subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 0 0.
Definition _set_Sedeleg_Fetch_Addr_Align
-(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
+(r_ref : register_ref regstate register_value Sedeleg) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6535,32 +5675,23 @@ Definition _set_Sedeleg_Fetch_Addr_Align
write_reg r_ref r
: M (unit).
-Definition _update_Sedeleg_Fetch_Addr_Align (v : Sedeleg) (x : mword 1)
-: Sedeleg :=
-
+Definition _update_Sedeleg_Fetch_Addr_Align (v : Sedeleg) (x : mword 1) : Sedeleg :=
{[ v with
Sedeleg_Sedeleg_chunk_0 :=
(update_subrange_vec_dec v.(Sedeleg_Sedeleg_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition legalize_sedeleg (s : Sedeleg) (v : mword 64)
-: Sedeleg :=
-
+Definition legalize_sedeleg (s : Sedeleg) (v : mword 64) : Sedeleg :=
Mk_Sedeleg (EXTZ 64 (subrange_vec_dec v 8 0)).
-Definition Mk_Sinterrupts (v : mword 64)
-: Sinterrupts :=
-
+Definition Mk_Sinterrupts (v : mword 64) : Sinterrupts :=
{| Sinterrupts_Sinterrupts_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Sinterrupts_bits (v : Sinterrupts)
-: mword 64 :=
-
+Definition _get_Sinterrupts_bits (v : Sinterrupts) : mword 64 :=
subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 63 0.
Definition _set_Sinterrupts_bits
-(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 64)
+(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6570,22 +5701,17 @@ Definition _set_Sinterrupts_bits
write_reg r_ref r
: M (unit).
-Definition _update_Sinterrupts_bits (v : Sinterrupts) (x : mword 64)
-: Sinterrupts :=
-
+Definition _update_Sinterrupts_bits (v : Sinterrupts) (x : mword 64) : Sinterrupts :=
{[ v with
Sinterrupts_Sinterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Sinterrupts_SEI (v : Sinterrupts)
-: mword 1 :=
-
+Definition _get_Sinterrupts_SEI (v : Sinterrupts) : mword 1 :=
subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 9 9.
Definition _set_Sinterrupts_SEI
-(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6595,22 +5721,17 @@ Definition _set_Sinterrupts_SEI
write_reg r_ref r
: M (unit).
-Definition _update_Sinterrupts_SEI (v : Sinterrupts) (x : mword 1)
-: Sinterrupts :=
-
+Definition _update_Sinterrupts_SEI (v : Sinterrupts) (x : mword 1) : Sinterrupts :=
{[ v with
Sinterrupts_Sinterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 9 9 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sinterrupts_UEI (v : Sinterrupts)
-: mword 1 :=
-
+Definition _get_Sinterrupts_UEI (v : Sinterrupts) : mword 1 :=
subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 8 8.
Definition _set_Sinterrupts_UEI
-(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6620,22 +5741,17 @@ Definition _set_Sinterrupts_UEI
write_reg r_ref r
: M (unit).
-Definition _update_Sinterrupts_UEI (v : Sinterrupts) (x : mword 1)
-: Sinterrupts :=
-
+Definition _update_Sinterrupts_UEI (v : Sinterrupts) (x : mword 1) : Sinterrupts :=
{[ v with
Sinterrupts_Sinterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sinterrupts_STI (v : Sinterrupts)
-: mword 1 :=
-
+Definition _get_Sinterrupts_STI (v : Sinterrupts) : mword 1 :=
subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 5 5.
Definition _set_Sinterrupts_STI
-(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6645,22 +5761,17 @@ Definition _set_Sinterrupts_STI
write_reg r_ref r
: M (unit).
-Definition _update_Sinterrupts_STI (v : Sinterrupts) (x : mword 1)
-: Sinterrupts :=
-
+Definition _update_Sinterrupts_STI (v : Sinterrupts) (x : mword 1) : Sinterrupts :=
{[ v with
Sinterrupts_Sinterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 5 5 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sinterrupts_UTI (v : Sinterrupts)
-: mword 1 :=
-
+Definition _get_Sinterrupts_UTI (v : Sinterrupts) : mword 1 :=
subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 4 4.
Definition _set_Sinterrupts_UTI
-(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6670,22 +5781,17 @@ Definition _set_Sinterrupts_UTI
write_reg r_ref r
: M (unit).
-Definition _update_Sinterrupts_UTI (v : Sinterrupts) (x : mword 1)
-: Sinterrupts :=
-
+Definition _update_Sinterrupts_UTI (v : Sinterrupts) (x : mword 1) : Sinterrupts :=
{[ v with
Sinterrupts_Sinterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sinterrupts_SSI (v : Sinterrupts)
-: mword 1 :=
-
+Definition _get_Sinterrupts_SSI (v : Sinterrupts) : mword 1 :=
subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 1 1.
Definition _set_Sinterrupts_SSI
-(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6695,22 +5801,17 @@ Definition _set_Sinterrupts_SSI
write_reg r_ref r
: M (unit).
-Definition _update_Sinterrupts_SSI (v : Sinterrupts) (x : mword 1)
-: Sinterrupts :=
-
+Definition _update_Sinterrupts_SSI (v : Sinterrupts) (x : mword 1) : Sinterrupts :=
{[ v with
Sinterrupts_Sinterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Sinterrupts_USI (v : Sinterrupts)
-: mword 1 :=
-
+Definition _get_Sinterrupts_USI (v : Sinterrupts) : mword 1 :=
subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 0 0.
Definition _set_Sinterrupts_USI
-(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
+(r_ref : register_ref regstate register_value Sinterrupts) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6720,17 +5821,13 @@ Definition _set_Sinterrupts_USI
write_reg r_ref r
: M (unit).
-Definition _update_Sinterrupts_USI (v : Sinterrupts) (x : mword 1)
-: Sinterrupts :=
-
+Definition _update_Sinterrupts_USI (v : Sinterrupts) (x : mword 1) : Sinterrupts :=
{[ v with
Sinterrupts_Sinterrupts_chunk_0 :=
(update_subrange_vec_dec v.(Sinterrupts_Sinterrupts_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition lower_mip (m : Minterrupts) (d : Minterrupts)
-: Sinterrupts :=
-
- let s : Sinterrupts := Mk_Sinterrupts (EXTZ 64 (vec_of_bits [B0] : mword 1)) in
+Definition lower_mip (m : Minterrupts) (d : Minterrupts) : Sinterrupts :=
+ let s : Sinterrupts := Mk_Sinterrupts (EXTZ 64 ('b"0" : mword 1)) in
let s := _update_Sinterrupts_SEI s (and_vec (_get_Minterrupts_SEI m) (_get_Minterrupts_SEI d)) in
let s := _update_Sinterrupts_STI s (and_vec (_get_Minterrupts_STI m) (_get_Minterrupts_STI d)) in
let s := _update_Sinterrupts_SSI s (and_vec (_get_Minterrupts_SSI m) (_get_Minterrupts_SSI d)) in
@@ -6738,10 +5835,8 @@ Definition lower_mip (m : Minterrupts) (d : Minterrupts)
let s := _update_Sinterrupts_UTI s (and_vec (_get_Minterrupts_UTI m) (_get_Minterrupts_UTI d)) in
_update_Sinterrupts_USI s (and_vec (_get_Minterrupts_USI m) (_get_Minterrupts_USI d)).
-Definition lower_mie (m : Minterrupts) (d : Minterrupts)
-: Sinterrupts :=
-
- let s : Sinterrupts := Mk_Sinterrupts (EXTZ 64 (vec_of_bits [B0] : mword 1)) in
+Definition lower_mie (m : Minterrupts) (d : Minterrupts) : Sinterrupts :=
+ let s : Sinterrupts := Mk_Sinterrupts (EXTZ 64 ('b"0" : mword 1)) in
let s := _update_Sinterrupts_SEI s (and_vec (_get_Minterrupts_SEI m) (_get_Minterrupts_SEI d)) in
let s := _update_Sinterrupts_STI s (and_vec (_get_Minterrupts_STI m) (_get_Minterrupts_STI d)) in
let s := _update_Sinterrupts_SSI s (and_vec (_get_Minterrupts_SSI m) (_get_Minterrupts_SSI d)) in
@@ -6749,80 +5844,63 @@ Definition lower_mie (m : Minterrupts) (d : Minterrupts)
let s := _update_Sinterrupts_UTI s (and_vec (_get_Minterrupts_UTI m) (_get_Minterrupts_UTI d)) in
_update_Sinterrupts_USI s (and_vec (_get_Minterrupts_USI m) (_get_Minterrupts_USI d)).
-Definition lift_sip (o : Minterrupts) (d : Minterrupts) (s : Sinterrupts)
-: M (Minterrupts) :=
-
+Definition lift_sip (o : Minterrupts) (d : Minterrupts) (s : Sinterrupts) : M (Minterrupts) :=
let m : Minterrupts := o in
let m := _update_Minterrupts_SSI m (and_vec (_get_Sinterrupts_SSI s) (_get_Minterrupts_SSI d)) in
(haveNExt tt) >>= fun w__0 : bool =>
- returnm ((if sumbool_of_bool (w__0) then
- let m :=
- if ((eq_vec (_get_Minterrupts_UEI d) ((bool_to_bits true) : mword 1))) then
- _update_Minterrupts_UEI m (_get_Sinterrupts_UEI s)
- else m in
- if ((eq_vec (_get_Minterrupts_USI d) ((bool_to_bits true) : mword 1))) then
- _update_Minterrupts_USI m (_get_Sinterrupts_USI s)
- else m
- else m)
- : Minterrupts).
-
-Definition legalize_sip (m : Minterrupts) (d : Minterrupts) (v : mword 64)
-: M (Minterrupts) :=
-
- (lift_sip m d (Mk_Sinterrupts v))
- : M (Minterrupts).
-
-Definition lift_sie (o : Minterrupts) (d : Minterrupts) (s : Sinterrupts)
-: M (Minterrupts) :=
-
+ returnm (if sumbool_of_bool w__0 then
+ let m :=
+ if eq_vec (_get_Minterrupts_UEI d) ('b"1" : mword 1) then
+ _update_Minterrupts_UEI m (_get_Sinterrupts_UEI s)
+ else m in
+ if eq_vec (_get_Minterrupts_USI d) ('b"1" : mword 1) then
+ _update_Minterrupts_USI m (_get_Sinterrupts_USI s)
+ else m
+ else m).
+
+Definition legalize_sip (m : Minterrupts) (d : Minterrupts) (v : mword 64) : M (Minterrupts) :=
+ (lift_sip m d (Mk_Sinterrupts v)) : M (Minterrupts).
+
+Definition lift_sie (o : Minterrupts) (d : Minterrupts) (s : Sinterrupts) : M (Minterrupts) :=
let m : Minterrupts := o in
let m :=
- if ((eq_vec (_get_Minterrupts_SEI d) ((bool_to_bits true) : mword 1))) then
+ if eq_vec (_get_Minterrupts_SEI d) ('b"1" : mword 1) then
_update_Minterrupts_SEI m (_get_Sinterrupts_SEI s)
else m in
let m :=
- if ((eq_vec (_get_Minterrupts_STI d) ((bool_to_bits true) : mword 1))) then
+ if eq_vec (_get_Minterrupts_STI d) ('b"1" : mword 1) then
_update_Minterrupts_STI m (_get_Sinterrupts_STI s)
else m in
let m :=
- if ((eq_vec (_get_Minterrupts_SSI d) ((bool_to_bits true) : mword 1))) then
+ if eq_vec (_get_Minterrupts_SSI d) ('b"1" : mword 1) then
_update_Minterrupts_SSI m (_get_Sinterrupts_SSI s)
else m in
(haveNExt tt) >>= fun w__0 : bool =>
- returnm ((if sumbool_of_bool (w__0) then
- let m :=
- if ((eq_vec (_get_Minterrupts_UEI d) ((bool_to_bits true) : mword 1))) then
- _update_Minterrupts_UEI m (_get_Sinterrupts_UEI s)
- else m in
- let m :=
- if ((eq_vec (_get_Minterrupts_UTI d) ((bool_to_bits true) : mword 1))) then
- _update_Minterrupts_UTI m (_get_Sinterrupts_UTI s)
- else m in
- if ((eq_vec (_get_Minterrupts_USI d) ((bool_to_bits true) : mword 1))) then
- _update_Minterrupts_USI m (_get_Sinterrupts_USI s)
- else m
- else m)
- : Minterrupts).
-
-Definition legalize_sie (m : Minterrupts) (d : Minterrupts) (v : mword 64)
-: M (Minterrupts) :=
-
- (lift_sie m d (Mk_Sinterrupts v))
- : M (Minterrupts).
-
-Definition Mk_Satp64 (v : mword 64)
-: Satp64 :=
-
+ returnm (if sumbool_of_bool w__0 then
+ let m :=
+ if eq_vec (_get_Minterrupts_UEI d) ('b"1" : mword 1) then
+ _update_Minterrupts_UEI m (_get_Sinterrupts_UEI s)
+ else m in
+ let m :=
+ if eq_vec (_get_Minterrupts_UTI d) ('b"1" : mword 1) then
+ _update_Minterrupts_UTI m (_get_Sinterrupts_UTI s)
+ else m in
+ if eq_vec (_get_Minterrupts_USI d) ('b"1" : mword 1) then
+ _update_Minterrupts_USI m (_get_Sinterrupts_USI s)
+ else m
+ else m).
+
+Definition legalize_sie (m : Minterrupts) (d : Minterrupts) (v : mword 64) : M (Minterrupts) :=
+ (lift_sie m d (Mk_Sinterrupts v)) : M (Minterrupts).
+
+Definition Mk_Satp64 (v : mword 64) : Satp64 :=
{| Satp64_Satp64_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_Satp64_bits (v : Satp64)
-: mword 64 :=
-
+Definition _get_Satp64_bits (v : Satp64) : mword 64 :=
subrange_vec_dec v.(Satp64_Satp64_chunk_0) 63 0.
-Definition _set_Satp64_bits (r_ref : register_ref regstate register_value Satp64) (v : mword 64)
+Definition _set_Satp64_bits (r_ref : register_ref regstate register_value Satp64) (v : mword 64)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6832,21 +5910,16 @@ Definition _set_Satp64_bits (r_ref : register_ref regstate register_value Satp64
write_reg r_ref r
: M (unit).
-Definition _update_Satp64_bits (v : Satp64) (x : mword 64)
-: Satp64 :=
-
+Definition _update_Satp64_bits (v : Satp64) (x : mword 64) : Satp64 :=
{[ v with
Satp64_Satp64_chunk_0 :=
(update_subrange_vec_dec v.(Satp64_Satp64_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_Satp64_Mode (v : Satp64)
-: mword 4 :=
-
+Definition _get_Satp64_Mode (v : Satp64) : mword 4 :=
subrange_vec_dec v.(Satp64_Satp64_chunk_0) 63 60.
-Definition _set_Satp64_Mode (r_ref : register_ref regstate register_value Satp64) (v : mword 4)
+Definition _set_Satp64_Mode (r_ref : register_ref regstate register_value Satp64) (v : mword 4)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6856,21 +5929,16 @@ Definition _set_Satp64_Mode (r_ref : register_ref regstate register_value Satp64
write_reg r_ref r
: M (unit).
-Definition _update_Satp64_Mode (v : Satp64) (x : mword 4)
-: Satp64 :=
-
+Definition _update_Satp64_Mode (v : Satp64) (x : mword 4) : Satp64 :=
{[ v with
Satp64_Satp64_chunk_0 :=
(update_subrange_vec_dec v.(Satp64_Satp64_chunk_0) 63 60 (subrange_vec_dec x 3 0)) ]}.
-Definition _get_Satp64_Asid (v : Satp64)
-: mword 16 :=
-
+Definition _get_Satp64_Asid (v : Satp64) : mword 16 :=
subrange_vec_dec v.(Satp64_Satp64_chunk_0) 59 44.
-Definition _set_Satp64_Asid (r_ref : register_ref regstate register_value Satp64) (v : mword 16)
+Definition _set_Satp64_Asid (r_ref : register_ref regstate register_value Satp64) (v : mword 16)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6880,21 +5948,16 @@ Definition _set_Satp64_Asid (r_ref : register_ref regstate register_value Satp64
write_reg r_ref r
: M (unit).
-Definition _update_Satp64_Asid (v : Satp64) (x : mword 16)
-: Satp64 :=
-
+Definition _update_Satp64_Asid (v : Satp64) (x : mword 16) : Satp64 :=
{[ v with
Satp64_Satp64_chunk_0 :=
(update_subrange_vec_dec v.(Satp64_Satp64_chunk_0) 59 44 (subrange_vec_dec x 15 0)) ]}.
-Definition _get_Satp64_PPN (v : Satp64)
-: mword 44 :=
-
+Definition _get_Satp64_PPN (v : Satp64) : mword 44 :=
subrange_vec_dec v.(Satp64_Satp64_chunk_0) 43 0.
-Definition _set_Satp64_PPN (r_ref : register_ref regstate register_value Satp64) (v : mword 44)
+Definition _set_Satp64_PPN (r_ref : register_ref regstate register_value Satp64) (v : mword 44)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6904,36 +5967,27 @@ Definition _set_Satp64_PPN (r_ref : register_ref regstate register_value Satp64)
write_reg r_ref r
: M (unit).
-Definition _update_Satp64_PPN (v : Satp64) (x : mword 44)
-: Satp64 :=
-
+Definition _update_Satp64_PPN (v : Satp64) (x : mword 44) : Satp64 :=
{[ v with
Satp64_Satp64_chunk_0 :=
(update_subrange_vec_dec v.(Satp64_Satp64_chunk_0) 43 0 (subrange_vec_dec x 43 0)) ]}.
-Definition legalize_satp64 (a : Architecture) (o : mword 64) (v : mword 64)
-: mword 64 :=
-
+Definition legalize_satp64 (a : Architecture) (o : mword 64) (v : mword 64) : mword 64 :=
let s := Mk_Satp64 v in
match (satp64Mode_of_bits a (_get_Satp64_Mode s)) with
| None => o
- | Some (Sv32) => o
- | Some (_) => _get_Satp64_bits s
+ | Some Sv32 => o
+ | Some _ => _get_Satp64_bits s
end.
-Definition Mk_Satp32 (v : mword 32)
-: Satp32 :=
-
+Definition Mk_Satp32 (v : mword 32) : Satp32 :=
{| Satp32_Satp32_chunk_0 := (subrange_vec_dec v 31 0) |}.
-Definition _get_Satp32_bits (v : Satp32)
-: mword 32 :=
-
+Definition _get_Satp32_bits (v : Satp32) : mword 32 :=
subrange_vec_dec v.(Satp32_Satp32_chunk_0) 31 0.
-Definition _set_Satp32_bits (r_ref : register_ref regstate register_value Satp32) (v : mword 32)
+Definition _set_Satp32_bits (r_ref : register_ref regstate register_value Satp32) (v : mword 32)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6943,21 +5997,16 @@ Definition _set_Satp32_bits (r_ref : register_ref regstate register_value Satp32
write_reg r_ref r
: M (unit).
-Definition _update_Satp32_bits (v : Satp32) (x : mword 32)
-: Satp32 :=
-
+Definition _update_Satp32_bits (v : Satp32) (x : mword 32) : Satp32 :=
{[ v with
Satp32_Satp32_chunk_0 :=
(update_subrange_vec_dec v.(Satp32_Satp32_chunk_0) 31 0 (subrange_vec_dec x 31 0)) ]}.
-Definition _get_Satp32_Mode (v : Satp32)
-: mword 1 :=
-
+Definition _get_Satp32_Mode (v : Satp32) : mword 1 :=
subrange_vec_dec v.(Satp32_Satp32_chunk_0) 31 31.
-Definition _set_Satp32_Mode (r_ref : register_ref regstate register_value Satp32) (v : mword 1)
+Definition _set_Satp32_Mode (r_ref : register_ref regstate register_value Satp32) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6967,21 +6016,16 @@ Definition _set_Satp32_Mode (r_ref : register_ref regstate register_value Satp32
write_reg r_ref r
: M (unit).
-Definition _update_Satp32_Mode (v : Satp32) (x : mword 1)
-: Satp32 :=
-
+Definition _update_Satp32_Mode (v : Satp32) (x : mword 1) : Satp32 :=
{[ v with
Satp32_Satp32_chunk_0 :=
(update_subrange_vec_dec v.(Satp32_Satp32_chunk_0) 31 31 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Satp32_Asid (v : Satp32)
-: mword 9 :=
-
+Definition _get_Satp32_Asid (v : Satp32) : mword 9 :=
subrange_vec_dec v.(Satp32_Satp32_chunk_0) 30 22.
-Definition _set_Satp32_Asid (r_ref : register_ref regstate register_value Satp32) (v : mword 9)
+Definition _set_Satp32_Asid (r_ref : register_ref regstate register_value Satp32) (v : mword 9)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -6991,21 +6035,16 @@ Definition _set_Satp32_Asid (r_ref : register_ref regstate register_value Satp32
write_reg r_ref r
: M (unit).
-Definition _update_Satp32_Asid (v : Satp32) (x : mword 9)
-: Satp32 :=
-
+Definition _update_Satp32_Asid (v : Satp32) (x : mword 9) : Satp32 :=
{[ v with
Satp32_Satp32_chunk_0 :=
(update_subrange_vec_dec v.(Satp32_Satp32_chunk_0) 30 22 (subrange_vec_dec x 8 0)) ]}.
-Definition _get_Satp32_PPN (v : Satp32)
-: mword 22 :=
-
+Definition _get_Satp32_PPN (v : Satp32) : mword 22 :=
subrange_vec_dec v.(Satp32_Satp32_chunk_0) 21 0.
-Definition _set_Satp32_PPN (r_ref : register_ref regstate register_value Satp32) (v : mword 22)
+Definition _set_Satp32_PPN (r_ref : register_ref regstate register_value Satp32) (v : mword 22)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -7015,68 +6054,53 @@ Definition _set_Satp32_PPN (r_ref : register_ref regstate register_value Satp32)
write_reg r_ref r
: M (unit).
-Definition _update_Satp32_PPN (v : Satp32) (x : mword 22)
-: Satp32 :=
-
+Definition _update_Satp32_PPN (v : Satp32) (x : mword 22) : Satp32 :=
{[ v with
Satp32_Satp32_chunk_0 :=
(update_subrange_vec_dec v.(Satp32_Satp32_chunk_0) 21 0 (subrange_vec_dec x 21 0)) ]}.
-Definition legalize_satp32 (a : Architecture) (o : mword 32) (v : mword 32) : mword 32 := v.
+Definition legalize_satp32 (a : Architecture) (o : mword 32) (v : mword 32) : mword 32 := v.
-Definition PmpAddrMatchType_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 3)}
+Definition PmpAddrMatchType_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 3))}
: PmpAddrMatchType :=
-
- let l__20 := arg_ in
- if sumbool_of_bool ((Z.eqb l__20 0)) then OFF
- else if sumbool_of_bool ((Z.eqb l__20 1)) then TOR
- else if sumbool_of_bool ((Z.eqb l__20 2)) then NA4
+ let l__82 := arg_ in
+ if sumbool_of_bool (Z.eqb l__82 0) then OFF
+ else if sumbool_of_bool (Z.eqb l__82 1) then TOR
+ else if sumbool_of_bool (Z.eqb l__82 2) then NA4
else NAPOT.
-Definition num_of_PmpAddrMatchType (arg_ : PmpAddrMatchType)
-: {e : Z & ArithFact (0 <= e /\ e <= 3)} :=
-
- build_ex(match arg_ with | OFF => 0 | TOR => 1 | NA4 => 2 | NAPOT => 3 end).
+Definition num_of_PmpAddrMatchType (arg_ : PmpAddrMatchType)
+: {e : Z & ArithFact ((0 <=? e) && (e <=? 3))} :=
+ build_ex (match arg_ with | OFF => 0 | TOR => 1 | NA4 => 2 | NAPOT => 3 end).
-Definition pmpAddrMatchType_of_bits (bs : mword 2)
-: M (PmpAddrMatchType) :=
-
+Definition pmpAddrMatchType_of_bits (bs : mword 2) : M (PmpAddrMatchType) :=
let b__0 := bs in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0] : mword 2))) then returnm (OFF : PmpAddrMatchType)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then returnm (TOR : PmpAddrMatchType)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then returnm (NA4 : PmpAddrMatchType)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then
- returnm (NAPOT
- : PmpAddrMatchType)
+ (if eq_vec b__0 ('b"00" : mword 2) then returnm OFF
+ else if eq_vec b__0 ('b"01" : mword 2) then returnm TOR
+ else if eq_vec b__0 ('b"10" : mword 2) then returnm NA4
+ else if eq_vec b__0 ('b"11" : mword 2) then returnm NAPOT
else
assert_exp' false "Pattern match failure at model/riscv_pmp_regs.sail 7:2 - 12:3" >>= fun _ =>
exit tt)
: M (PmpAddrMatchType).
-Definition pmpAddrMatchType_to_bits (bs : PmpAddrMatchType)
-: mword 2 :=
-
+Definition pmpAddrMatchType_to_bits (bs : PmpAddrMatchType) : mword 2 :=
match bs with
- | OFF => (vec_of_bits [B0;B0] : mword 2)
- | TOR => (vec_of_bits [B0;B1] : mword 2)
- | NA4 => (vec_of_bits [B1;B0] : mword 2)
- | NAPOT => (vec_of_bits [B1;B1] : mword 2)
+ | OFF => 'b"00" : mword 2
+ | TOR => 'b"01" : mword 2
+ | NA4 => 'b"10" : mword 2
+ | NAPOT => 'b"11" : mword 2
end.
-Definition Mk_Pmpcfg_ent (v : mword 8)
-: Pmpcfg_ent :=
-
+Definition Mk_Pmpcfg_ent (v : mword 8) : Pmpcfg_ent :=
{| Pmpcfg_ent_Pmpcfg_ent_chunk_0 := (subrange_vec_dec v 7 0) |}.
-Definition _get_Pmpcfg_ent_bits (v : Pmpcfg_ent)
-: mword 8 :=
-
+Definition _get_Pmpcfg_ent_bits (v : Pmpcfg_ent) : mword 8 :=
subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 7 0.
Definition _set_Pmpcfg_ent_bits
-(r_ref : register_ref regstate register_value Pmpcfg_ent) (v : mword 8)
+(r_ref : register_ref regstate register_value Pmpcfg_ent) (v : mword 8)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -7086,22 +6110,16 @@ Definition _set_Pmpcfg_ent_bits
write_reg r_ref r
: M (unit).
-Definition _update_Pmpcfg_ent_bits (v : Pmpcfg_ent) (x : mword 8)
-: Pmpcfg_ent :=
-
+Definition _update_Pmpcfg_ent_bits (v : Pmpcfg_ent) (x : mword 8) : Pmpcfg_ent :=
{[ v with
Pmpcfg_ent_Pmpcfg_ent_chunk_0 :=
(update_subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 7 0 (subrange_vec_dec x 7 0)) ]}.
-Definition _get_Pmpcfg_ent_L (v : Pmpcfg_ent)
-: mword 1 :=
-
+Definition _get_Pmpcfg_ent_L (v : Pmpcfg_ent) : mword 1 :=
subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 7 7.
Definition _set_Pmpcfg_ent_L (r_ref : register_ref regstate register_value Pmpcfg_ent) (v : mword 1)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -7111,22 +6129,16 @@ Definition _set_Pmpcfg_ent_L (r_ref : register_ref regstate register_value Pmpcf
write_reg r_ref r
: M (unit).
-Definition _update_Pmpcfg_ent_L (v : Pmpcfg_ent) (x : mword 1)
-: Pmpcfg_ent :=
-
+Definition _update_Pmpcfg_ent_L (v : Pmpcfg_ent) (x : mword 1) : Pmpcfg_ent :=
{[ v with
Pmpcfg_ent_Pmpcfg_ent_chunk_0 :=
(update_subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 7 7 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Pmpcfg_ent_A (v : Pmpcfg_ent)
-: mword 2 :=
-
+Definition _get_Pmpcfg_ent_A (v : Pmpcfg_ent) : mword 2 :=
subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 4 3.
Definition _set_Pmpcfg_ent_A (r_ref : register_ref regstate register_value Pmpcfg_ent) (v : mword 2)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -7136,22 +6148,16 @@ Definition _set_Pmpcfg_ent_A (r_ref : register_ref regstate register_value Pmpcf
write_reg r_ref r
: M (unit).
-Definition _update_Pmpcfg_ent_A (v : Pmpcfg_ent) (x : mword 2)
-: Pmpcfg_ent :=
-
+Definition _update_Pmpcfg_ent_A (v : Pmpcfg_ent) (x : mword 2) : Pmpcfg_ent :=
{[ v with
Pmpcfg_ent_Pmpcfg_ent_chunk_0 :=
(update_subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 4 3 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_Pmpcfg_ent_X (v : Pmpcfg_ent)
-: mword 1 :=
-
+Definition _get_Pmpcfg_ent_X (v : Pmpcfg_ent) : mword 1 :=
subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 2 2.
Definition _set_Pmpcfg_ent_X (r_ref : register_ref regstate register_value Pmpcfg_ent) (v : mword 1)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -7161,22 +6167,16 @@ Definition _set_Pmpcfg_ent_X (r_ref : register_ref regstate register_value Pmpcf
write_reg r_ref r
: M (unit).
-Definition _update_Pmpcfg_ent_X (v : Pmpcfg_ent) (x : mword 1)
-: Pmpcfg_ent :=
-
+Definition _update_Pmpcfg_ent_X (v : Pmpcfg_ent) (x : mword 1) : Pmpcfg_ent :=
{[ v with
Pmpcfg_ent_Pmpcfg_ent_chunk_0 :=
(update_subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 2 2 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Pmpcfg_ent_W (v : Pmpcfg_ent)
-: mword 1 :=
-
+Definition _get_Pmpcfg_ent_W (v : Pmpcfg_ent) : mword 1 :=
subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 1 1.
Definition _set_Pmpcfg_ent_W (r_ref : register_ref regstate register_value Pmpcfg_ent) (v : mword 1)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -7186,22 +6186,16 @@ Definition _set_Pmpcfg_ent_W (r_ref : register_ref regstate register_value Pmpcf
write_reg r_ref r
: M (unit).
-Definition _update_Pmpcfg_ent_W (v : Pmpcfg_ent) (x : mword 1)
-: Pmpcfg_ent :=
-
+Definition _update_Pmpcfg_ent_W (v : Pmpcfg_ent) (x : mword 1) : Pmpcfg_ent :=
{[ v with
Pmpcfg_ent_Pmpcfg_ent_chunk_0 :=
(update_subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_Pmpcfg_ent_R (v : Pmpcfg_ent)
-: mword 1 :=
-
+Definition _get_Pmpcfg_ent_R (v : Pmpcfg_ent) : mword 1 :=
subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 0 0.
Definition _set_Pmpcfg_ent_R (r_ref : register_ref regstate register_value Pmpcfg_ent) (v : mword 1)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -7211,18 +6205,14 @@ Definition _set_Pmpcfg_ent_R (r_ref : register_ref regstate register_value Pmpcf
write_reg r_ref r
: M (unit).
-Definition _update_Pmpcfg_ent_R (v : Pmpcfg_ent) (x : mword 1)
-: Pmpcfg_ent :=
-
+Definition _update_Pmpcfg_ent_R (v : Pmpcfg_ent) (x : mword 1) : Pmpcfg_ent :=
{[ v with
Pmpcfg_ent_Pmpcfg_ent_chunk_0 :=
(update_subrange_vec_dec v.(Pmpcfg_ent_Pmpcfg_ent_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition pmpReadCfgReg (n : Z) `{ArithFact (0 <= n /\ n < 4)}
-: M (mword 64) :=
-
- let l__18 := n in
- (if sumbool_of_bool ((Z.eqb l__18 0)) then
+Definition pmpReadCfgReg (n : Z) `{ArithFact ((0 <=? n) && (n <? 4))} : M (mword 64) :=
+ let l__80 := n in
+ (if sumbool_of_bool (Z.eqb l__80 0) then
read_reg pmp7cfg_ref >>= fun w__0 : Pmpcfg_ent =>
read_reg pmp6cfg_ref >>= fun w__1 : Pmpcfg_ent =>
read_reg pmp5cfg_ref >>= fun w__2 : Pmpcfg_ent =>
@@ -7231,15 +6221,14 @@ Definition pmpReadCfgReg (n : Z) `{ArithFact (0 <= n /\ n < 4)}
read_reg pmp2cfg_ref >>= fun w__5 : Pmpcfg_ent =>
read_reg pmp1cfg_ref >>= fun w__6 : Pmpcfg_ent =>
read_reg pmp0cfg_ref >>= fun w__7 : Pmpcfg_ent =>
- returnm ((concat_vec (_get_Pmpcfg_ent_bits w__0)
- (concat_vec (_get_Pmpcfg_ent_bits w__1)
- (concat_vec (_get_Pmpcfg_ent_bits w__2)
- (concat_vec (_get_Pmpcfg_ent_bits w__3)
- (concat_vec (_get_Pmpcfg_ent_bits w__4)
- (concat_vec (_get_Pmpcfg_ent_bits w__5)
- (concat_vec (_get_Pmpcfg_ent_bits w__6) (_get_Pmpcfg_ent_bits w__7))))))))
- : mword (8 + (8 + (8 + (8 + (8 + (8 + (8 + 8))))))))
- else if sumbool_of_bool ((Z.eqb l__18 2)) then
+ returnm (concat_vec (_get_Pmpcfg_ent_bits w__0)
+ (concat_vec (_get_Pmpcfg_ent_bits w__1)
+ (concat_vec (_get_Pmpcfg_ent_bits w__2)
+ (concat_vec (_get_Pmpcfg_ent_bits w__3)
+ (concat_vec (_get_Pmpcfg_ent_bits w__4)
+ (concat_vec (_get_Pmpcfg_ent_bits w__5)
+ (concat_vec (_get_Pmpcfg_ent_bits w__6) (_get_Pmpcfg_ent_bits w__7))))))))
+ else if sumbool_of_bool (Z.eqb l__80 2) then
read_reg pmp15cfg_ref >>= fun w__8 : Pmpcfg_ent =>
read_reg pmp14cfg_ref >>= fun w__9 : Pmpcfg_ent =>
read_reg pmp13cfg_ref >>= fun w__10 : Pmpcfg_ent =>
@@ -7248,31 +6237,33 @@ Definition pmpReadCfgReg (n : Z) `{ArithFact (0 <= n /\ n < 4)}
read_reg pmp10cfg_ref >>= fun w__13 : Pmpcfg_ent =>
read_reg pmp9cfg_ref >>= fun w__14 : Pmpcfg_ent =>
read_reg pmp8cfg_ref >>= fun w__15 : Pmpcfg_ent =>
- returnm ((concat_vec (_get_Pmpcfg_ent_bits w__8)
- (concat_vec (_get_Pmpcfg_ent_bits w__9)
- (concat_vec (_get_Pmpcfg_ent_bits w__10)
- (concat_vec (_get_Pmpcfg_ent_bits w__11)
- (concat_vec (_get_Pmpcfg_ent_bits w__12)
- (concat_vec (_get_Pmpcfg_ent_bits w__13)
- (concat_vec (_get_Pmpcfg_ent_bits w__14)
- (_get_Pmpcfg_ent_bits w__15))))))))
- : mword (8 + (8 + (8 + (8 + (8 + (8 + (8 + 8))))))))
+ returnm (concat_vec (_get_Pmpcfg_ent_bits w__8)
+ (concat_vec (_get_Pmpcfg_ent_bits w__9)
+ (concat_vec (_get_Pmpcfg_ent_bits w__10)
+ (concat_vec (_get_Pmpcfg_ent_bits w__11)
+ (concat_vec (_get_Pmpcfg_ent_bits w__12)
+ (concat_vec (_get_Pmpcfg_ent_bits w__13)
+ (concat_vec (_get_Pmpcfg_ent_bits w__14)
+ (_get_Pmpcfg_ent_bits w__15))))))))
else
assert_exp' false "Pattern match failure at model/riscv_pmp_regs.sail 75:2 - 85:8" >>= fun _ =>
exit tt)
: M (mword 64).
-Definition pmpWriteCfg (cfg : Pmpcfg_ent) (v : mword 8)
-: Pmpcfg_ent :=
-
- if ((eq_vec (_get_Pmpcfg_ent_L cfg) ((bool_to_bits true) : mword 1))) then cfg
- else Mk_Pmpcfg_ent v.
+Definition pmpLocked (cfg : Pmpcfg_ent) : bool := eq_vec (_get_Pmpcfg_ent_L cfg) ('b"1" : mword 1).
-Definition pmpWriteCfgReg (n : Z) (v : mword 64) `{ArithFact (0 <= n /\ n < 4)}
-: M (unit) :=
-
- let l__16 := n in
- (if sumbool_of_bool ((Z.eqb l__16 0)) then
+Definition pmpTORLocked (cfg : Pmpcfg_ent) : M (bool) :=
+ (and_boolM (returnm ((eq_vec (_get_Pmpcfg_ent_L cfg) ('b"1" : mword 1)) : bool))
+ ((pmpAddrMatchType_of_bits (_get_Pmpcfg_ent_A cfg)) >>= fun w__0 : PmpAddrMatchType =>
+ returnm ((generic_eq w__0 TOR) : bool)))
+ : M (bool).
+
+Definition pmpWriteCfg (cfg : Pmpcfg_ent) (v : mword 8) : Pmpcfg_ent :=
+ if pmpLocked cfg then cfg else Mk_Pmpcfg_ent (and_vec v (Ox"9F" : mword 8)).
+
+Definition pmpWriteCfgReg (n : Z) (v : mword 64) `{ArithFact ((0 <=? n) && (n <? 4))} : M (unit) :=
+ let l__78 := n in
+ (if sumbool_of_bool (Z.eqb l__78 0) then
read_reg pmp0cfg_ref >>= fun w__0 : Pmpcfg_ent =>
write_reg pmp0cfg_ref (pmpWriteCfg w__0 (subrange_vec_dec v 7 0)) >>
read_reg pmp1cfg_ref >>= fun w__1 : Pmpcfg_ent =>
@@ -7290,11 +6281,11 @@ Definition pmpWriteCfgReg (n : Z) (v : mword 64) `{ArithFact (0 <= n /\ n < 4)}
read_reg pmp7cfg_ref >>= fun w__7 : Pmpcfg_ent =>
write_reg pmp7cfg_ref (pmpWriteCfg w__7 (subrange_vec_dec v 63 56))
: M (unit)
- else if sumbool_of_bool ((Z.eqb l__16 2)) then
+ else if sumbool_of_bool (Z.eqb l__78 2) then
read_reg pmp8cfg_ref >>= fun w__8 : Pmpcfg_ent =>
- let pmp8cfg8 := pmpWriteCfg w__8 (subrange_vec_dec v 7 0) in
+ write_reg pmp8cfg_ref (pmpWriteCfg w__8 (subrange_vec_dec v 7 0)) >>
read_reg pmp9cfg_ref >>= fun w__9 : Pmpcfg_ent =>
- let pmp9cfg9 := pmpWriteCfg w__9 (subrange_vec_dec v 15 8) in
+ write_reg pmp9cfg_ref (pmpWriteCfg w__9 (subrange_vec_dec v 15 8)) >>
read_reg pmp10cfg_ref >>= fun w__10 : Pmpcfg_ent =>
write_reg pmp10cfg_ref (pmpWriteCfg w__10 (subrange_vec_dec v 23 16)) >>
read_reg pmp11cfg_ref >>= fun w__11 : Pmpcfg_ent =>
@@ -7309,280 +6300,184 @@ Definition pmpWriteCfgReg (n : Z) (v : mword 64) `{ArithFact (0 <= n /\ n < 4)}
write_reg pmp15cfg_ref (pmpWriteCfg w__15 (subrange_vec_dec v 63 56))
: M (unit)
else
- assert_exp' false "Pattern match failure at model/riscv_pmp_regs.sail 94:2 - 137:8" >>= fun _ =>
+ assert_exp' false "Pattern match failure at model/riscv_pmp_regs.sail 101:2 - 144:8" >>= fun _ =>
exit tt)
: M (unit).
-Definition pmpWriteAddr (cfg : Pmpcfg_ent) (reg : mword 64) (v : mword 64)
+Definition pmpWriteAddr (locked : bool) (tor_locked : bool) (reg : mword 64) (v : mword 64)
: mword 64 :=
-
- if ((eq_vec (_get_Pmpcfg_ent_L cfg) ((bool_to_bits true) : mword 1))) then reg
- else EXTZ 64 (subrange_vec_dec v 53 0).
+ if sumbool_of_bool (orb locked tor_locked) then reg else EXTZ 64 (subrange_vec_dec v 53 0).
-Definition pmpAddrRange (cfg : Pmpcfg_ent) (pmpaddr : mword 64) (prev_pmpaddr : mword 64)
+Definition pmpAddrRange (cfg : Pmpcfg_ent) (pmpaddr : mword 64) (prev_pmpaddr : mword 64)
: M (option ((mword 64 * mword 64))) :=
-
(pmpAddrMatchType_of_bits (_get_Pmpcfg_ent_A cfg)) >>= fun w__0 : PmpAddrMatchType =>
- returnm ((match w__0 with
- | OFF => None
- | TOR => Some ((shiftl prev_pmpaddr 2, shiftl pmpaddr 2))
- | NA4 =>
- let lo := shiftl pmpaddr 2 in
- Some
- ((lo, add_vec_int lo 4))
- | NAPOT =>
- let mask := xor_vec pmpaddr (add_vec_int pmpaddr 1) in
- let lo := and_vec pmpaddr (not_vec mask) in
- let len := add_vec_int mask 1 in
- Some
- ((shiftl lo 2, shiftl (add_vec lo len) 2))
- end)
- : option ((mword 64 * mword 64))).
+ returnm (match w__0 with
+ | OFF => None
+ | TOR => Some (shiftl prev_pmpaddr 2, shiftl pmpaddr 2)
+ | NA4 =>
+ let lo := shiftl pmpaddr 2 in
+ Some (lo, add_vec_int lo 4)
+ | NAPOT =>
+ let mask := xor_vec pmpaddr (add_vec_int pmpaddr 1) in
+ let lo := and_vec pmpaddr (not_vec mask) in
+ let len := add_vec_int mask 1 in
+ Some (shiftl lo 2, shiftl (add_vec lo len) 2)
+ end).
-Definition pmpCheckRWX (ent : Pmpcfg_ent) (acc : AccessType)
-: bool :=
-
+Definition pmpCheckRWX (ent : Pmpcfg_ent) (acc : AccessType unit) : bool :=
match acc with
- | Read => eq_vec (_get_Pmpcfg_ent_R ent) ((bool_to_bits true) : mword 1)
- | Write => eq_vec (_get_Pmpcfg_ent_W ent) ((bool_to_bits true) : mword 1)
- | ReadWrite =>
- andb (eq_vec (_get_Pmpcfg_ent_R ent) ((bool_to_bits true) : mword 1))
- (eq_vec (_get_Pmpcfg_ent_W ent) ((bool_to_bits true) : mword 1))
- | Execute => eq_vec (_get_Pmpcfg_ent_X ent) ((bool_to_bits true) : mword 1)
+ | Read _ => eq_vec (_get_Pmpcfg_ent_R ent) ('b"1" : mword 1)
+ | Write _ => eq_vec (_get_Pmpcfg_ent_W ent) ('b"1" : mword 1)
+ | ReadWrite _ =>
+ andb (eq_vec (_get_Pmpcfg_ent_R ent) ('b"1" : mword 1))
+ (eq_vec (_get_Pmpcfg_ent_W ent) ('b"1" : mword 1))
+ | Execute tt => eq_vec (_get_Pmpcfg_ent_X ent) ('b"1" : mword 1)
end.
-Definition pmpCheckPerms (ent : Pmpcfg_ent) (acc : AccessType) (priv : Privilege)
-: bool :=
-
+Definition pmpCheckPerms (ent : Pmpcfg_ent) (acc : AccessType unit) (priv : Privilege) : bool :=
match priv with
- | Machine =>
- if ((eq_vec (_get_Pmpcfg_ent_L ent) ((bool_to_bits true) : mword 1))) then
- pmpCheckRWX ent acc
- else true
+ | Machine => if pmpLocked ent then pmpCheckRWX ent acc else true
| _ => pmpCheckRWX ent acc
end.
-Definition pmpAddrMatch_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
-: pmpAddrMatch :=
-
- let l__14 := arg_ in
- if sumbool_of_bool ((Z.eqb l__14 0)) then PMP_NoMatch
- else if sumbool_of_bool ((Z.eqb l__14 1)) then PMP_PartialMatch
+Definition pmpAddrMatch_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : pmpAddrMatch :=
+ let l__76 := arg_ in
+ if sumbool_of_bool (Z.eqb l__76 0) then PMP_NoMatch
+ else if sumbool_of_bool (Z.eqb l__76 1) then PMP_PartialMatch
else PMP_Match.
-Definition num_of_pmpAddrMatch (arg_ : pmpAddrMatch)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | PMP_NoMatch => 0 | PMP_PartialMatch => 1 | PMP_Match => 2 end).
+Definition num_of_pmpAddrMatch (arg_ : pmpAddrMatch) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | PMP_NoMatch => 0 | PMP_PartialMatch => 1 | PMP_Match => 2 end).
-Definition pmpMatchAddr (addr : mword 64) (width : mword 64) (rng : option ((mword 64 * mword 64)))
+Definition pmpMatchAddr (addr : mword 64) (width : mword 64) (rng : option ((mword 64 * mword 64)))
: pmpAddrMatch :=
-
match rng with
| None => PMP_NoMatch
- | Some ((lo, hi)) =>
- if ((zopz0zI_u hi lo)) then PMP_NoMatch
- else if ((orb (zopz0zI_u (add_vec addr width) lo) (zopz0zI_u hi addr))) then PMP_NoMatch
- else if ((andb (zopz0zIzJ_u lo addr) (zopz0zIzJ_u (add_vec addr width) hi))) then PMP_Match
+ | Some (lo, hi) =>
+ if zopz0zI_u hi lo then PMP_NoMatch
+ else if orb (zopz0zI_u (add_vec addr width) lo) (zopz0zI_u hi addr) then PMP_NoMatch
+ else if andb (zopz0zIzJ_u lo addr) (zopz0zIzJ_u (add_vec addr width) hi) then PMP_Match
else PMP_PartialMatch
end.
-Definition pmpMatch_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 2)}
-: pmpMatch :=
-
- let l__12 := arg_ in
- if sumbool_of_bool ((Z.eqb l__12 0)) then PMP_Success
- else if sumbool_of_bool ((Z.eqb l__12 1)) then PMP_Continue
+Definition pmpMatch_of_num (arg_ : Z) `{ArithFact ((0 <=? arg_) && (arg_ <=? 2))} : pmpMatch :=
+ let l__74 := arg_ in
+ if sumbool_of_bool (Z.eqb l__74 0) then PMP_Success
+ else if sumbool_of_bool (Z.eqb l__74 1) then PMP_Continue
else PMP_Fail.
-Definition num_of_pmpMatch (arg_ : pmpMatch)
-: {e : Z & ArithFact (0 <= e /\ e <= 2)} :=
-
- build_ex(match arg_ with | PMP_Success => 0 | PMP_Continue => 1 | PMP_Fail => 2 end).
+Definition num_of_pmpMatch (arg_ : pmpMatch) : {e : Z & ArithFact ((0 <=? e) && (e <=? 2))} :=
+ build_ex (match arg_ with | PMP_Success => 0 | PMP_Continue => 1 | PMP_Fail => 2 end).
Definition pmpMatchEntry
-(addr : mword 64) (width : mword 64) (acc : AccessType) (priv : Privilege) (ent : Pmpcfg_ent)
-(pmpaddr : mword 64) (prev_pmpaddr : mword 64)
+(addr : mword 64) (width : mword 64) (acc : AccessType unit) (priv : Privilege) (ent : Pmpcfg_ent)
+(pmpaddr : mword 64) (prev_pmpaddr : mword 64)
: M (pmpMatch) :=
-
(pmpAddrRange ent pmpaddr prev_pmpaddr) >>= fun rng =>
- returnm ((match (pmpMatchAddr addr width rng) with
- | PMP_NoMatch => PMP_Continue
- | PMP_PartialMatch => PMP_Fail
- | PMP_Match => if ((pmpCheckPerms ent acc priv)) then PMP_Success else PMP_Fail
- end)
- : pmpMatch).
+ returnm (match (pmpMatchAddr addr width rng) with
+ | PMP_NoMatch => PMP_Continue
+ | PMP_PartialMatch => PMP_Fail
+ | PMP_Match => if pmpCheckPerms ent acc priv then PMP_Success else PMP_Fail
+ end).
-Definition pmpCheck (addr : mword 64) (width : Z) (acc : AccessType) (priv : Privilege)
-`{ArithFact (width > 0)}
+Definition pmpCheck (addr : mword 64) (width : Z) (acc : AccessType unit) (priv : Privilege)
+`{ArithFact (width >? 0)}
: M (option ExceptionType) :=
-
let width : xlenbits := to_bits 64 width in
read_reg pmp0cfg_ref >>= fun w__0 : Pmpcfg_ent =>
((read_reg pmpaddr0_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
(pmpMatchEntry addr width acc priv w__0 w__1 (zeros_implicit 64)) >>= fun w__2 : pmpMatch =>
(match w__2 with
- | PMP_Success => returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp1cfg_ref >>= fun w__3 : Pmpcfg_ent =>
((read_reg pmpaddr1_ref) : M (mword 64)) >>= fun w__4 : mword 64 =>
((read_reg pmpaddr0_ref) : M (mword 64)) >>= fun w__5 : mword 64 =>
(pmpMatchEntry addr width acc priv w__3 w__4 w__5) >>= fun w__6 : pmpMatch =>
(match w__6 with
- | PMP_Success =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp2cfg_ref >>= fun w__7 : Pmpcfg_ent =>
((read_reg pmpaddr2_ref) : M (mword 64)) >>= fun w__8 : mword 64 =>
((read_reg pmpaddr1_ref) : M (mword 64)) >>= fun w__9 : mword 64 =>
(pmpMatchEntry addr width acc priv w__7 w__8 w__9) >>= fun w__10 : pmpMatch =>
(match w__10 with
- | PMP_Success =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp3cfg_ref >>= fun w__11 : Pmpcfg_ent =>
((read_reg pmpaddr3_ref) : M (mword 64)) >>= fun w__12 : mword 64 =>
((read_reg pmpaddr2_ref) : M (mword 64)) >>= fun w__13 : mword 64 =>
(pmpMatchEntry addr width acc priv w__11 w__12 w__13) >>= fun w__14 : pmpMatch =>
(match w__14 with
- | PMP_Success =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp4cfg_ref >>= fun w__15 : Pmpcfg_ent =>
((read_reg pmpaddr4_ref) : M (mword 64)) >>= fun w__16 : mword 64 =>
((read_reg pmpaddr3_ref) : M (mword 64)) >>= fun w__17 : mword 64 =>
(pmpMatchEntry addr width acc priv w__15 w__16 w__17) >>= fun w__18 : pmpMatch =>
(match w__18 with
- | PMP_Success =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp5cfg_ref >>= fun w__19 : Pmpcfg_ent =>
((read_reg pmpaddr5_ref) : M (mword 64)) >>= fun w__20 : mword 64 =>
((read_reg pmpaddr4_ref) : M (mword 64)) >>= fun w__21 : mword 64 =>
(pmpMatchEntry addr width acc priv w__19 w__20 w__21) >>= fun w__22 : pmpMatch =>
(match w__22 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp6cfg_ref >>= fun w__23 : Pmpcfg_ent =>
((read_reg pmpaddr6_ref) : M (mword 64)) >>= fun w__24 : mword 64 =>
((read_reg pmpaddr5_ref) : M (mword 64)) >>= fun w__25 : mword 64 =>
(pmpMatchEntry addr width acc priv w__23 w__24 w__25) >>= fun w__26 : pmpMatch =>
(match w__26 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp7cfg_ref >>= fun w__27 : Pmpcfg_ent =>
((read_reg pmpaddr7_ref) : M (mword 64)) >>= fun w__28 : mword 64 =>
((read_reg pmpaddr6_ref) : M (mword 64)) >>= fun w__29 : mword 64 =>
(pmpMatchEntry addr width acc priv w__27 w__28 w__29) >>= fun w__30 : pmpMatch =>
(match w__30 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp8cfg_ref >>= fun w__31 : Pmpcfg_ent =>
((read_reg pmpaddr8_ref) : M (mword 64)) >>= fun w__32 : mword 64 =>
((read_reg pmpaddr7_ref) : M (mword 64)) >>= fun w__33 : mword 64 =>
(pmpMatchEntry addr width acc priv w__31 w__32 w__33) >>= fun w__34 : pmpMatch =>
(match w__34 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp9cfg_ref >>= fun w__35 : Pmpcfg_ent =>
((read_reg pmpaddr9_ref) : M (mword 64)) >>= fun w__36 : mword 64 =>
((read_reg pmpaddr8_ref) : M (mword 64)) >>= fun w__37 : mword 64 =>
(pmpMatchEntry addr width acc priv w__35 w__36 w__37) >>= fun w__38 : pmpMatch =>
(match w__38 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp10cfg_ref >>= fun w__39 : Pmpcfg_ent =>
((read_reg pmpaddr10_ref) : M (mword 64)) >>= fun w__40 : mword 64 =>
((read_reg pmpaddr9_ref) : M (mword 64)) >>= fun w__41 : mword 64 =>
(pmpMatchEntry addr width acc priv w__39 w__40 w__41) >>= fun w__42 : pmpMatch =>
(match w__42 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp11cfg_ref >>= fun w__43 : Pmpcfg_ent =>
((read_reg pmpaddr11_ref) : M (mword 64)) >>= fun w__44 : mword 64 =>
((read_reg pmpaddr10_ref) : M (mword 64)) >>= fun w__45 : mword 64 =>
(pmpMatchEntry addr width acc priv w__43 w__44 w__45) >>= fun w__46 : pmpMatch =>
(match w__46 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp12cfg_ref >>= fun w__47 : Pmpcfg_ent =>
((read_reg pmpaddr12_ref) : M (mword 64)) >>= fun w__48 : mword 64 =>
@@ -7590,17 +6485,8 @@ Definition pmpCheck (addr : mword 64) (width : Z) (acc : AccessType) (priv : Pri
(pmpMatchEntry addr width acc priv w__47 w__48
w__49) >>= fun w__50 : pmpMatch =>
(match w__50 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp13cfg_ref >>= fun w__51 : Pmpcfg_ent =>
((read_reg pmpaddr13_ref) : M (mword 64)) >>= fun w__52 : mword 64 =>
@@ -7608,18 +6494,8 @@ Definition pmpCheck (addr : mword 64) (width : Z) (acc : AccessType) (priv : Pri
(pmpMatchEntry addr width acc priv w__51
w__52 w__53) >>= fun w__54 : pmpMatch =>
(match w__54 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp14cfg_ref >>= fun w__55 : Pmpcfg_ent =>
((read_reg pmpaddr14_ref) : M (mword 64)) >>= fun w__56 : mword 64 =>
@@ -7627,18 +6503,8 @@ Definition pmpCheck (addr : mword 64) (width : Z) (acc : AccessType) (priv : Pri
(pmpMatchEntry addr width acc priv w__55
w__56 w__57) >>= fun w__58 : pmpMatch =>
(match w__58 with
- | PMP_Success =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
- | PMP_Fail =>
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))}))
+ | PMP_Success => returnm true
+ | PMP_Fail => returnm false
| PMP_Continue =>
read_reg pmp15cfg_ref >>= fun w__59 : Pmpcfg_ent =>
((read_reg pmpaddr15_ref)
@@ -7647,16 +6513,15 @@ Definition pmpCheck (addr : mword 64) (width : Z) (acc : AccessType) (priv : Pri
: M (mword 64)) >>= fun w__61 : mword 64 =>
(pmpMatchEntry addr width acc priv
w__59 w__60 w__61) >>= fun w__62 : pmpMatch =>
- returnm ((match w__62 with
- | PMP_Success => true
- | PMP_Fail => false
- | PMP_Continue =>
- match priv with
- | Machine => true
- | _ => false
- end
- end)
- : bool)
+ returnm (match w__62 with
+ | PMP_Success => true
+ | PMP_Fail => false
+ | PMP_Continue =>
+ match priv with
+ | Machine => true
+ | _ => false
+ end
+ end)
end)
: M (bool)
end)
@@ -7686,19 +6551,16 @@ Definition pmpCheck (addr : mword 64) (width : Z) (acc : AccessType) (priv : Pri
end)
: M (bool)
end) >>= fun check' : bool =>
- returnm ((if sumbool_of_bool (check') then None
- else
- match acc with
- | Read => Some (E_Load_Access_Fault)
- | Write => Some (E_SAMO_Access_Fault)
- | ReadWrite => Some (E_SAMO_Access_Fault)
- | Execute => Some (E_Fetch_Access_Fault)
- end)
- : option ExceptionType).
-
-Definition init_pmp '(tt : unit)
-: M (unit) :=
-
+ returnm (if sumbool_of_bool check' then None
+ else
+ match acc with
+ | Read _ => Some (E_Load_Access_Fault tt)
+ | Write _ => Some (E_SAMO_Access_Fault tt)
+ | ReadWrite _ => Some (E_SAMO_Access_Fault tt)
+ | Execute tt => Some (E_Fetch_Access_Fault tt)
+ end).
+
+Definition init_pmp '(tt : unit) : M (unit) :=
read_reg pmp0cfg_ref >>= fun w__0 : Pmpcfg_ent =>
write_reg pmp0cfg_ref (_update_Pmpcfg_ent_A w__0 (pmpAddrMatchType_to_bits OFF)) >>
read_reg pmp1cfg_ref >>= fun w__1 : Pmpcfg_ent =>
@@ -7733,2584 +6595,3897 @@ Definition init_pmp '(tt : unit)
write_reg pmp15cfg_ref (_update_Pmpcfg_ent_A w__15 (pmpAddrMatchType_to_bits OFF))
: M (unit).
-Definition ext_init_regs '(tt : unit) : M (unit) := returnm (tt : unit).
+Definition ext_init_regs '(tt : unit) : M (unit) := returnm tt.
+
+Definition ext_rvfi_init '(tt : unit) : M (unit) :=
+ ((read_reg x1_ref) : M (mword 64)) >>= fun w__0 : mword 64 => write_reg x1_ref w__0 : M (unit).
-Definition ext_rvfi_init '(tt : unit) : M (unit) := returnm (tt : unit).
+Definition ext_check_CSR (csrno : mword 12) (p : Privilege) (isWrite : bool) : bool := true.
-Definition ext_fetch_check_pc (start_pc : mword 64) (pc : mword 64)
-: Ext_FetchAddr_Check unit :=
-
- Ext_FetchAddr_OK
- (pc).
+Definition ext_check_CSR_fail '(tt : unit) : unit := tt.
-Definition ext_handle_fetch_check_error (err : unit) : unit := tt.
+Definition ext_fetch_check_pc (start_pc : mword 64) (pc : mword 64) : Ext_FetchAddr_Check unit :=
+ Ext_FetchAddr_OK pc.
-Definition ext_control_check_addr (pc : mword 64)
-: Ext_ControlAddr_Check unit :=
-
- Ext_ControlAddr_OK
- (pc).
+Definition ext_handle_fetch_check_error (err : unit) : unit := tt.
-Definition ext_control_check_pc (pc : mword 64)
-: Ext_ControlAddr_Check unit :=
-
- Ext_ControlAddr_OK
- (pc).
+Definition ext_control_check_addr (pc : mword 64) : Ext_ControlAddr_Check unit :=
+ Ext_ControlAddr_OK pc.
-Definition ext_handle_control_check_error (err : unit) : unit := tt.
+Definition ext_control_check_pc (pc : mword 64) : Ext_ControlAddr_Check unit :=
+ Ext_ControlAddr_OK pc.
+
+Definition ext_handle_control_check_error (err : unit) : unit := tt.
Definition ext_data_get_addr
-(base : mword 5) (offset : mword 64) (acc : AccessType) (width : word_width)
+(base : mword 5) (offset : mword 64) (acc : AccessType unit) (width : word_width)
: M (Ext_DataAddr_Check unit) :=
-
- (rX (projT1 (regidx_to_regno base))) >>= fun w__0 : mword 64 =>
+ (rX_bits base) >>= fun w__0 : mword 64 =>
let addr := add_vec w__0 offset in
- returnm ((Ext_DataAddr_OK
- (addr))
- : Ext_DataAddr_Check unit).
+ returnm (Ext_DataAddr_OK addr).
-Definition ext_handle_data_check_error (err : unit) : unit := tt.
+Definition ext_handle_data_check_error (err : unit) : unit := tt.
-Definition csr_name (csr : mword 12)
-: string :=
-
- let b__0 := csr in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then "ustatus"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- "uie"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- "utvec"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "uscratch"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- "uepc"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "ucause"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- "utval"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- "uip"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- "fflags"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "frm"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- "fcsr"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "cycle"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- "time"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "instret"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "cycleh"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- "timeh"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "instreth"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "sstatus"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "sedeleg"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- "sideleg"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- "sie"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- "stvec"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
- "scounteren"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "sscratch"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- "sepc"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "scause"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- "stval"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- "sip"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "satp"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then
- "mvendorid"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then
- "marchid"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then
- "mimpid"
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then
- "mhartid"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "mstatus"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- "misa"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "medeleg"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- "mideleg"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- "mie"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- "mtvec"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
- "mcounteren"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "mscratch"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- "mepc"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "mcause"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- "mtval"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- "mip"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- "pmpcfg0"
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then
- "pmpaddr0"
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "mcycle"
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "minstret"
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- "mcycleh"
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- "minstreth"
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- "tselect"
- else "UNKNOWN".
-
-Definition csr_name_map_forwards (arg_ : mword 12)
-: M (string) :=
-
- let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("ustatus"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- returnm ("uie"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- returnm ("utvec"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("uscratch"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("uepc"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("ucause"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ("utval"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- returnm ("uip"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("fflags"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("frm"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ("fcsr"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("cycle"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("time"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("instret"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("cycleh"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("timeh"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("instreth"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("sstatus"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("sedeleg"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ("sideleg"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- returnm ("sie"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- returnm ("stvec"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
- returnm ("scounteren"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("sscratch"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("sepc"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("scause"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ("stval"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- returnm ("sip"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("satp"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then
- returnm ("mvendorid"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then
- returnm ("marchid"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then
- returnm ("mimpid"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then
- returnm ("mhartid"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("mstatus"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("misa"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("medeleg"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ("mideleg"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- returnm ("mie"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- returnm ("mtvec"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
- returnm ("mcounteren"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("mscratch"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("mepc"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("mcause"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ("mtval"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- returnm ("mip"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("pmpcfg0"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("pmpcfg1"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("pmpcfg2"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ("pmpcfg3"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then
- returnm ("pmpaddr0"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then
- returnm ("pmpaddr1"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then
- returnm ("pmpaddr2"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then
- returnm ("pmpaddr3"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then
- returnm ("pmpaddr4"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then
- returnm ("pmpaddr5"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then
- returnm ("pmpaddr6"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then
- returnm ("pmpaddr7"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then
- returnm ("pmpaddr8"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then
- returnm ("pmpaddr9"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then
- returnm ("pmpaddr10"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then
- returnm ("pmpaddr11"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then
- returnm ("pmpaddr12"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then
- returnm ("pmpaddr13"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then
- returnm ("pmpaddr14"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then
- returnm ("pmpaddr15"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("mcycle"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("minstret"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("mcycleh"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("minstreth"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ("tselect"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ("tdata1"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ("tdata2"
- : string)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ("tdata3"
- : string)
- else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
- : M (string).
+Definition Mk_Ustatus (v : mword 64) : Ustatus :=
+ {| Ustatus_Ustatus_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition csr_name_map_backwards (arg_ : string)
-: M (mword 12) :=
-
- let p0_ := arg_ in
- (if ((generic_eq p0_ "ustatus")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "uie")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "utvec")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "uscratch")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "uepc")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "ucause")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "utval")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "uip")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "fflags")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "frm")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "fcsr")) then
- returnm ((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "cycle")) then
- returnm ((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "time")) then
- returnm ((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "instret")) then
- returnm ((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "cycleh")) then
- returnm ((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "timeh")) then
- returnm ((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "instreth")) then
- returnm ((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "sstatus")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "sedeleg")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "sideleg")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "sie")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "stvec")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "scounteren")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "sscratch")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "sepc")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "scause")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "stval")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "sip")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "satp")) then
- returnm ((vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mvendorid")) then
- returnm ((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "marchid")) then
- returnm ((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mimpid")) then
- returnm ((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mhartid")) then
- returnm ((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mstatus")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "misa")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "medeleg")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mideleg")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mie")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mtvec")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mcounteren")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mscratch")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mepc")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mcause")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mtval")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mip")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpcfg0")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpcfg1")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpcfg2")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpcfg3")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr0")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr1")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr2")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr3")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr4")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr5")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr6")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr7")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr8")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr9")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr10")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr11")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr12")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr13")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr14")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "pmpaddr15")) then
- returnm ((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mcycle")) then
- returnm ((vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "minstret")) then
- returnm ((vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "mcycleh")) then
- returnm ((vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "minstreth")) then
- returnm ((vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "tselect")) then
- returnm ((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "tdata1")) then
- returnm ((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "tdata2")) then
- returnm ((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12)
- : mword 12)
- else if ((generic_eq p0_ "tdata3")) then
- returnm ((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12)
- : mword 12)
- else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
- : M (mword 12).
+Definition _get_Ustatus_bits (v : Ustatus) : mword 64 :=
+ subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 63 0.
-Definition csr_name_map_forwards_matches (arg_ : mword 12)
-: bool :=
-
- let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then true
- else false.
+Definition _set_Ustatus_bits (r_ref : register_ref regstate register_value Ustatus) (v : mword 64)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Ustatus_Ustatus_chunk_0 :=
+ (update_subrange_vec_dec r.(Ustatus_Ustatus_chunk_0) 63 0 (subrange_vec_dec v 63 0)) ]}
+ : Ustatus in
+ write_reg r_ref r
+ : M (unit).
-Definition csr_name_map_backwards_matches (arg_ : string)
-: bool :=
-
- let p0_ := arg_ in
- if ((generic_eq p0_ "ustatus")) then true
- else if ((generic_eq p0_ "uie")) then true
- else if ((generic_eq p0_ "utvec")) then true
- else if ((generic_eq p0_ "uscratch")) then true
- else if ((generic_eq p0_ "uepc")) then true
- else if ((generic_eq p0_ "ucause")) then true
- else if ((generic_eq p0_ "utval")) then true
- else if ((generic_eq p0_ "uip")) then true
- else if ((generic_eq p0_ "fflags")) then true
- else if ((generic_eq p0_ "frm")) then true
- else if ((generic_eq p0_ "fcsr")) then true
- else if ((generic_eq p0_ "cycle")) then true
- else if ((generic_eq p0_ "time")) then true
- else if ((generic_eq p0_ "instret")) then true
- else if ((generic_eq p0_ "cycleh")) then true
- else if ((generic_eq p0_ "timeh")) then true
- else if ((generic_eq p0_ "instreth")) then true
- else if ((generic_eq p0_ "sstatus")) then true
- else if ((generic_eq p0_ "sedeleg")) then true
- else if ((generic_eq p0_ "sideleg")) then true
- else if ((generic_eq p0_ "sie")) then true
- else if ((generic_eq p0_ "stvec")) then true
- else if ((generic_eq p0_ "scounteren")) then true
- else if ((generic_eq p0_ "sscratch")) then true
- else if ((generic_eq p0_ "sepc")) then true
- else if ((generic_eq p0_ "scause")) then true
- else if ((generic_eq p0_ "stval")) then true
- else if ((generic_eq p0_ "sip")) then true
- else if ((generic_eq p0_ "satp")) then true
- else if ((generic_eq p0_ "mvendorid")) then true
- else if ((generic_eq p0_ "marchid")) then true
- else if ((generic_eq p0_ "mimpid")) then true
- else if ((generic_eq p0_ "mhartid")) then true
- else if ((generic_eq p0_ "mstatus")) then true
- else if ((generic_eq p0_ "misa")) then true
- else if ((generic_eq p0_ "medeleg")) then true
- else if ((generic_eq p0_ "mideleg")) then true
- else if ((generic_eq p0_ "mie")) then true
- else if ((generic_eq p0_ "mtvec")) then true
- else if ((generic_eq p0_ "mcounteren")) then true
- else if ((generic_eq p0_ "mscratch")) then true
- else if ((generic_eq p0_ "mepc")) then true
- else if ((generic_eq p0_ "mcause")) then true
- else if ((generic_eq p0_ "mtval")) then true
- else if ((generic_eq p0_ "mip")) then true
- else if ((generic_eq p0_ "pmpcfg0")) then true
- else if ((generic_eq p0_ "pmpcfg1")) then true
- else if ((generic_eq p0_ "pmpcfg2")) then true
- else if ((generic_eq p0_ "pmpcfg3")) then true
- else if ((generic_eq p0_ "pmpaddr0")) then true
- else if ((generic_eq p0_ "pmpaddr1")) then true
- else if ((generic_eq p0_ "pmpaddr2")) then true
- else if ((generic_eq p0_ "pmpaddr3")) then true
- else if ((generic_eq p0_ "pmpaddr4")) then true
- else if ((generic_eq p0_ "pmpaddr5")) then true
- else if ((generic_eq p0_ "pmpaddr6")) then true
- else if ((generic_eq p0_ "pmpaddr7")) then true
- else if ((generic_eq p0_ "pmpaddr8")) then true
- else if ((generic_eq p0_ "pmpaddr9")) then true
- else if ((generic_eq p0_ "pmpaddr10")) then true
- else if ((generic_eq p0_ "pmpaddr11")) then true
- else if ((generic_eq p0_ "pmpaddr12")) then true
- else if ((generic_eq p0_ "pmpaddr13")) then true
- else if ((generic_eq p0_ "pmpaddr14")) then true
- else if ((generic_eq p0_ "pmpaddr15")) then true
- else if ((generic_eq p0_ "mcycle")) then true
- else if ((generic_eq p0_ "minstret")) then true
- else if ((generic_eq p0_ "mcycleh")) then true
- else if ((generic_eq p0_ "minstreth")) then true
- else if ((generic_eq p0_ "tselect")) then true
- else if ((generic_eq p0_ "tdata1")) then true
- else if ((generic_eq p0_ "tdata2")) then true
- else if ((generic_eq p0_ "tdata3")) then true
- else false.
+Definition _update_Ustatus_bits (v : Ustatus) (x : mword 64) : Ustatus :=
+ {[ v with
+ Ustatus_Ustatus_chunk_0 :=
+ (update_subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _s488_ (_s489_ : string)
-: option string :=
-
- let _s490_ := _s489_ in
- if ((string_startswith _s490_ "tdata3")) then
- match (string_drop _s490_ (projT1 (string_length "tdata3"))) with | s_ => Some (s_) end
- else None.
+Definition _get_Ustatus_UPIE (v : Ustatus) : mword 1 :=
+ subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 4 4.
-Definition _s484_ (_s485_ : string)
-: option string :=
-
- let _s486_ := _s485_ in
- if ((string_startswith _s486_ "tdata2")) then
- match (string_drop _s486_ (projT1 (string_length "tdata2"))) with | s_ => Some (s_) end
- else None.
+Definition _set_Ustatus_UPIE (r_ref : register_ref regstate register_value Ustatus) (v : mword 1)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Ustatus_Ustatus_chunk_0 :=
+ (update_subrange_vec_dec r.(Ustatus_Ustatus_chunk_0) 4 4 (subrange_vec_dec v 0 0)) ]}
+ : Ustatus in
+ write_reg r_ref r
+ : M (unit).
-Definition _s480_ (_s481_ : string)
-: option string :=
-
- let _s482_ := _s481_ in
- if ((string_startswith _s482_ "tdata1")) then
- match (string_drop _s482_ (projT1 (string_length "tdata1"))) with | s_ => Some (s_) end
- else None.
+Definition _update_Ustatus_UPIE (v : Ustatus) (x : mword 1) : Ustatus :=
+ {[ v with
+ Ustatus_Ustatus_chunk_0 :=
+ (update_subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _s476_ (_s477_ : string)
-: option string :=
-
- let _s478_ := _s477_ in
- if ((string_startswith _s478_ "tselect")) then
- match (string_drop _s478_ (projT1 (string_length "tselect"))) with | s_ => Some (s_) end
- else None.
+Definition _get_Ustatus_UIE (v : Ustatus) : mword 1 :=
+ subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 0 0.
-Definition _s472_ (_s473_ : string)
-: option string :=
-
- let _s474_ := _s473_ in
- if ((string_startswith _s474_ "minstreth")) then
- match (string_drop _s474_ (projT1 (string_length "minstreth"))) with | s_ => Some (s_) end
- else None.
+Definition _set_Ustatus_UIE (r_ref : register_ref regstate register_value Ustatus) (v : mword 1)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Ustatus_Ustatus_chunk_0 :=
+ (update_subrange_vec_dec r.(Ustatus_Ustatus_chunk_0) 0 0 (subrange_vec_dec v 0 0)) ]}
+ : Ustatus in
+ write_reg r_ref r
+ : M (unit).
-Definition _s468_ (_s469_ : string)
-: option string :=
-
- let _s470_ := _s469_ in
- if ((string_startswith _s470_ "mcycleh")) then
- match (string_drop _s470_ (projT1 (string_length "mcycleh"))) with | s_ => Some (s_) end
- else None.
+Definition _update_Ustatus_UIE (v : Ustatus) (x : mword 1) : Ustatus :=
+ {[ v with
+ Ustatus_Ustatus_chunk_0 :=
+ (update_subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition _s464_ (_s465_ : string)
-: option string :=
-
- let _s466_ := _s465_ in
- if ((string_startswith _s466_ "minstret")) then
- match (string_drop _s466_ (projT1 (string_length "minstret"))) with | s_ => Some (s_) end
- else None.
+Definition lower_sstatus (s : Sstatus) : Ustatus :=
+ let u := Mk_Ustatus (EXTZ 64 ('b"0" : mword 1)) in
+ let u := _update_Ustatus_UPIE u (_get_Sstatus_UPIE s) in
+ _update_Ustatus_UIE u (_get_Sstatus_UIE s).
-Definition _s460_ (_s461_ : string)
-: option string :=
-
- let _s462_ := _s461_ in
- if ((string_startswith _s462_ "mcycle")) then
- match (string_drop _s462_ (projT1 (string_length "mcycle"))) with | s_ => Some (s_) end
- else None.
+Definition lift_ustatus (s : Sstatus) (u : Ustatus) : Sstatus :=
+ let s := _update_Sstatus_UPIE s (_get_Ustatus_UPIE u) in
+ _update_Sstatus_UIE s (_get_Ustatus_UIE u).
-Definition _s456_ (_s457_ : string)
-: option string :=
-
- let _s458_ := _s457_ in
- if ((string_startswith _s458_ "pmpaddr15")) then
- match (string_drop _s458_ (projT1 (string_length "pmpaddr15"))) with | s_ => Some (s_) end
- else None.
+Definition legalize_ustatus (m : Mstatus) (v : mword 64) : M (Mstatus) :=
+ let u := Mk_Ustatus v in
+ let s := lower_mstatus m in
+ let s := lift_ustatus s u in
+ lift_sstatus m s.
-Definition _s452_ (_s453_ : string)
-: option string :=
-
- let _s454_ := _s453_ in
- if ((string_startswith _s454_ "pmpaddr14")) then
- match (string_drop _s454_ (projT1 (string_length "pmpaddr14"))) with | s_ => Some (s_) end
- else None.
+Definition Mk_Uinterrupts (v : mword 64) : Uinterrupts :=
+ {| Uinterrupts_Uinterrupts_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _s448_ (_s449_ : string)
-: option string :=
-
- let _s450_ := _s449_ in
- if ((string_startswith _s450_ "pmpaddr13")) then
- match (string_drop _s450_ (projT1 (string_length "pmpaddr13"))) with | s_ => Some (s_) end
- else None.
+Definition _get_Uinterrupts_bits (v : Uinterrupts) : mword 64 :=
+ subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 63 0.
-Definition _s444_ (_s445_ : string)
-: option string :=
-
- let _s446_ := _s445_ in
- if ((string_startswith _s446_ "pmpaddr12")) then
- match (string_drop _s446_ (projT1 (string_length "pmpaddr12"))) with | s_ => Some (s_) end
- else None.
+Definition _set_Uinterrupts_bits
+(r_ref : register_ref regstate register_value Uinterrupts) (v : mword 64)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Uinterrupts_Uinterrupts_chunk_0 :=
+ (update_subrange_vec_dec r.(Uinterrupts_Uinterrupts_chunk_0) 63 0 (subrange_vec_dec v 63 0)) ]}
+ : Uinterrupts in
+ write_reg r_ref r
+ : M (unit).
-Definition _s440_ (_s441_ : string)
-: option string :=
-
- let _s442_ := _s441_ in
- if ((string_startswith _s442_ "pmpaddr11")) then
- match (string_drop _s442_ (projT1 (string_length "pmpaddr11"))) with | s_ => Some (s_) end
- else None.
+Definition _update_Uinterrupts_bits (v : Uinterrupts) (x : mword 64) : Uinterrupts :=
+ {[ v with
+ Uinterrupts_Uinterrupts_chunk_0 :=
+ (update_subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _s436_ (_s437_ : string)
-: option string :=
-
- let _s438_ := _s437_ in
- if ((string_startswith _s438_ "pmpaddr10")) then
- match (string_drop _s438_ (projT1 (string_length "pmpaddr10"))) with | s_ => Some (s_) end
- else None.
+Definition _get_Uinterrupts_UEI (v : Uinterrupts) : mword 1 :=
+ subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 8 8.
-Definition _s432_ (_s433_ : string)
-: option string :=
-
- let _s434_ := _s433_ in
- if ((string_startswith _s434_ "pmpaddr9")) then
- match (string_drop _s434_ (projT1 (string_length "pmpaddr9"))) with | s_ => Some (s_) end
- else None.
+Definition _set_Uinterrupts_UEI
+(r_ref : register_ref regstate register_value Uinterrupts) (v : mword 1)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Uinterrupts_Uinterrupts_chunk_0 :=
+ (update_subrange_vec_dec r.(Uinterrupts_Uinterrupts_chunk_0) 8 8 (subrange_vec_dec v 0 0)) ]}
+ : Uinterrupts in
+ write_reg r_ref r
+ : M (unit).
-Definition _s428_ (_s429_ : string)
-: option string :=
-
- let _s430_ := _s429_ in
- if ((string_startswith _s430_ "pmpaddr8")) then
- match (string_drop _s430_ (projT1 (string_length "pmpaddr8"))) with | s_ => Some (s_) end
- else None.
+Definition _update_Uinterrupts_UEI (v : Uinterrupts) (x : mword 1) : Uinterrupts :=
+ {[ v with
+ Uinterrupts_Uinterrupts_chunk_0 :=
+ (update_subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
-Definition _s424_ (_s425_ : string)
-: option string :=
-
- let _s426_ := _s425_ in
- if ((string_startswith _s426_ "pmpaddr7")) then
- match (string_drop _s426_ (projT1 (string_length "pmpaddr7"))) with | s_ => Some (s_) end
- else None.
+Definition _get_Uinterrupts_UTI (v : Uinterrupts) : mword 1 :=
+ subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 4 4.
-Definition _s420_ (_s421_ : string)
-: option string :=
-
- let _s422_ := _s421_ in
- if ((string_startswith _s422_ "pmpaddr6")) then
- match (string_drop _s422_ (projT1 (string_length "pmpaddr6"))) with | s_ => Some (s_) end
- else None.
+Definition _set_Uinterrupts_UTI
+(r_ref : register_ref regstate register_value Uinterrupts) (v : mword 1)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Uinterrupts_Uinterrupts_chunk_0 :=
+ (update_subrange_vec_dec r.(Uinterrupts_Uinterrupts_chunk_0) 4 4 (subrange_vec_dec v 0 0)) ]}
+ : Uinterrupts in
+ write_reg r_ref r
+ : M (unit).
-Definition _s416_ (_s417_ : string)
-: option string :=
-
- let _s418_ := _s417_ in
- if ((string_startswith _s418_ "pmpaddr5")) then
- match (string_drop _s418_ (projT1 (string_length "pmpaddr5"))) with | s_ => Some (s_) end
- else None.
+Definition _update_Uinterrupts_UTI (v : Uinterrupts) (x : mword 1) : Uinterrupts :=
+ {[ v with
+ Uinterrupts_Uinterrupts_chunk_0 :=
+ (update_subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _s412_ (_s413_ : string)
-: option string :=
-
- let _s414_ := _s413_ in
- if ((string_startswith _s414_ "pmpaddr4")) then
- match (string_drop _s414_ (projT1 (string_length "pmpaddr4"))) with | s_ => Some (s_) end
- else None.
+Definition _get_Uinterrupts_USI (v : Uinterrupts) : mword 1 :=
+ subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 0 0.
-Definition _s408_ (_s409_ : string)
-: option string :=
-
- let _s410_ := _s409_ in
- if ((string_startswith _s410_ "pmpaddr3")) then
- match (string_drop _s410_ (projT1 (string_length "pmpaddr3"))) with | s_ => Some (s_) end
- else None.
+Definition _set_Uinterrupts_USI
+(r_ref : register_ref regstate register_value Uinterrupts) (v : mword 1)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ Uinterrupts_Uinterrupts_chunk_0 :=
+ (update_subrange_vec_dec r.(Uinterrupts_Uinterrupts_chunk_0) 0 0 (subrange_vec_dec v 0 0)) ]}
+ : Uinterrupts in
+ write_reg r_ref r
+ : M (unit).
-Definition _s404_ (_s405_ : string)
-: option string :=
-
- let _s406_ := _s405_ in
- if ((string_startswith _s406_ "pmpaddr2")) then
- match (string_drop _s406_ (projT1 (string_length "pmpaddr2"))) with | s_ => Some (s_) end
- else None.
+Definition _update_Uinterrupts_USI (v : Uinterrupts) (x : mword 1) : Uinterrupts :=
+ {[ v with
+ Uinterrupts_Uinterrupts_chunk_0 :=
+ (update_subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition _s400_ (_s401_ : string)
-: option string :=
-
- let _s402_ := _s401_ in
- if ((string_startswith _s402_ "pmpaddr1")) then
- match (string_drop _s402_ (projT1 (string_length "pmpaddr1"))) with | s_ => Some (s_) end
- else None.
+Definition lower_sip (s : Sinterrupts) (d : Sinterrupts) : Uinterrupts :=
+ let u : Uinterrupts := Mk_Uinterrupts (EXTZ 64 ('b"0" : mword 1)) in
+ let u := _update_Uinterrupts_UEI u (and_vec (_get_Sinterrupts_UEI s) (_get_Sinterrupts_UEI d)) in
+ let u := _update_Uinterrupts_UTI u (and_vec (_get_Sinterrupts_UTI s) (_get_Sinterrupts_UTI d)) in
+ _update_Uinterrupts_USI u (and_vec (_get_Sinterrupts_USI s) (_get_Sinterrupts_USI d)).
-Definition _s396_ (_s397_ : string)
-: option string :=
-
- let _s398_ := _s397_ in
- if ((string_startswith _s398_ "pmpaddr0")) then
- match (string_drop _s398_ (projT1 (string_length "pmpaddr0"))) with | s_ => Some (s_) end
- else None.
+Definition lower_sie (s : Sinterrupts) (d : Sinterrupts) : Uinterrupts :=
+ let u : Uinterrupts := Mk_Uinterrupts (EXTZ 64 ('b"0" : mword 1)) in
+ let u := _update_Uinterrupts_UEI u (and_vec (_get_Sinterrupts_UEI s) (_get_Sinterrupts_UEI d)) in
+ let u := _update_Uinterrupts_UTI u (and_vec (_get_Sinterrupts_UTI s) (_get_Sinterrupts_UTI d)) in
+ _update_Uinterrupts_USI u (and_vec (_get_Sinterrupts_USI s) (_get_Sinterrupts_USI d)).
-Definition _s392_ (_s393_ : string)
-: option string :=
-
- let _s394_ := _s393_ in
- if ((string_startswith _s394_ "pmpcfg3")) then
- match (string_drop _s394_ (projT1 (string_length "pmpcfg3"))) with | s_ => Some (s_) end
- else None.
+Definition lift_uip (o : Sinterrupts) (d : Sinterrupts) (u : Uinterrupts) : Sinterrupts :=
+ let s : Sinterrupts := o in
+ if eq_vec (_get_Sinterrupts_USI d) ('b"1" : mword 1) then
+ _update_Sinterrupts_USI s (_get_Uinterrupts_USI u)
+ else s.
-Definition _s388_ (_s389_ : string)
-: option string :=
-
- let _s390_ := _s389_ in
- if ((string_startswith _s390_ "pmpcfg2")) then
- match (string_drop _s390_ (projT1 (string_length "pmpcfg2"))) with | s_ => Some (s_) end
- else None.
+Definition legalize_uip (s : Sinterrupts) (d : Sinterrupts) (v : mword 64) : Sinterrupts :=
+ lift_uip s d (Mk_Uinterrupts v).
-Definition _s384_ (_s385_ : string)
-: option string :=
-
- let _s386_ := _s385_ in
- if ((string_startswith _s386_ "pmpcfg1")) then
- match (string_drop _s386_ (projT1 (string_length "pmpcfg1"))) with | s_ => Some (s_) end
- else None.
+Definition lift_uie (o : Sinterrupts) (d : Sinterrupts) (u : Uinterrupts) : Sinterrupts :=
+ let s : Sinterrupts := o in
+ let s :=
+ if eq_vec (_get_Sinterrupts_UEI d) ('b"1" : mword 1) then
+ _update_Sinterrupts_UEI s (_get_Uinterrupts_UEI u)
+ else s in
+ let s :=
+ if eq_vec (_get_Sinterrupts_UTI d) ('b"1" : mword 1) then
+ _update_Sinterrupts_UTI s (_get_Uinterrupts_UTI u)
+ else s in
+ if eq_vec (_get_Sinterrupts_USI d) ('b"1" : mword 1) then
+ _update_Sinterrupts_USI s (_get_Uinterrupts_USI u)
+ else s.
-Definition _s380_ (_s381_ : string)
-: option string :=
-
- let _s382_ := _s381_ in
- if ((string_startswith _s382_ "pmpcfg0")) then
- match (string_drop _s382_ (projT1 (string_length "pmpcfg0"))) with | s_ => Some (s_) end
- else None.
+Definition legalize_uie (s : Sinterrupts) (d : Sinterrupts) (v : mword 64) : Sinterrupts :=
+ lift_uie s d (Mk_Uinterrupts v).
-Definition _s376_ (_s377_ : string)
-: option string :=
-
- let _s378_ := _s377_ in
- if ((string_startswith _s378_ "mip")) then
- match (string_drop _s378_ (projT1 (string_length "mip"))) with | s_ => Some (s_) end
- else None.
+Definition ext_check_xret_priv (p : Privilege) : bool := true.
-Definition _s372_ (_s373_ : string)
-: option string :=
-
- let _s374_ := _s373_ in
- if ((string_startswith _s374_ "mtval")) then
- match (string_drop _s374_ (projT1 (string_length "mtval"))) with | s_ => Some (s_) end
- else None.
+Definition ext_fail_xret_priv '(tt : unit) : unit := tt.
-Definition _s368_ (_s369_ : string)
-: option string :=
-
- let _s370_ := _s369_ in
- if ((string_startswith _s370_ "mcause")) then
- match (string_drop _s370_ (projT1 (string_length "mcause"))) with | s_ => Some (s_) end
- else None.
+Definition handle_trap_extension (p : Privilege) (pc : mword 64) (u : option unit) : unit := tt.
-Definition _s364_ (_s365_ : string)
-: option string :=
-
- let _s366_ := _s365_ in
- if ((string_startswith _s366_ "mepc")) then
- match (string_drop _s366_ (projT1 (string_length "mepc"))) with | s_ => Some (s_) end
- else None.
+Definition prepare_trap_vector (p : Privilege) (cause : Mcause) : M (mword 64) :=
+ (match p with
+ | Machine => read_reg mtvec_ref : M (Mtvec)
+ | Supervisor => read_reg stvec_ref : M (Mtvec)
+ | User => read_reg utvec_ref : M (Mtvec)
+ end) >>= fun tvec : Mtvec =>
+ (match (tvec_addr tvec cause) with
+ | Some epc => returnm epc
+ | None => (internal_error "Invalid tvec mode") : M (mword 64)
+ end)
+ : M (mword 64).
-Definition _s360_ (_s361_ : string)
-: option string :=
-
- let _s362_ := _s361_ in
- if ((string_startswith _s362_ "mscratch")) then
- match (string_drop _s362_ (projT1 (string_length "mscratch"))) with | s_ => Some (s_) end
- else None.
+Definition get_xret_target (p : Privilege) : M (mword 64) :=
+ (match p with
+ | Machine => ((read_reg mepc_ref) : M (mword 64)) : M (mword 64)
+ | Supervisor => ((read_reg sepc_ref) : M (mword 64)) : M (mword 64)
+ | User => ((read_reg uepc_ref) : M (mword 64)) : M (mword 64)
+ end)
+ : M (mword 64).
-Definition _s356_ (_s357_ : string)
-: option string :=
-
- let _s358_ := _s357_ in
- if ((string_startswith _s358_ "mcounteren")) then
- match (string_drop _s358_ (projT1 (string_length "mcounteren"))) with | s_ => Some (s_) end
- else None.
+Definition set_xret_target (p : Privilege) (value : mword 64) : M (mword 64) :=
+ (legalize_xepc value) >>= fun target =>
+ (match p with
+ | Machine => write_reg mepc_ref target : M (unit)
+ | Supervisor => write_reg sepc_ref target : M (unit)
+ | User => write_reg uepc_ref target : M (unit)
+ end) >>
+ returnm target.
-Definition _s352_ (_s353_ : string)
-: option string :=
-
- let _s354_ := _s353_ in
- if ((string_startswith _s354_ "mtvec")) then
- match (string_drop _s354_ (projT1 (string_length "mtvec"))) with | s_ => Some (s_) end
- else None.
+Definition prepare_xret_target (p : Privilege) : M (mword 64) :=
+ (get_xret_target p) : M (mword 64).
-Definition _s348_ (_s349_ : string)
-: option string :=
-
- let _s350_ := _s349_ in
- if ((string_startswith _s350_ "mie")) then
- match (string_drop _s350_ (projT1 (string_length "mie"))) with | s_ => Some (s_) end
- else None.
+Definition get_mtvec '(tt : unit) : M (mword 64) :=
+ read_reg mtvec_ref >>= fun w__0 : Mtvec => returnm (_get_Mtvec_bits w__0).
-Definition _s344_ (_s345_ : string)
-: option string :=
-
- let _s346_ := _s345_ in
- if ((string_startswith _s346_ "mideleg")) then
- match (string_drop _s346_ (projT1 (string_length "mideleg"))) with | s_ => Some (s_) end
- else None.
+Definition get_stvec '(tt : unit) : M (mword 64) :=
+ read_reg stvec_ref >>= fun w__0 : Mtvec => returnm (_get_Mtvec_bits w__0).
-Definition _s340_ (_s341_ : string)
-: option string :=
-
- let _s342_ := _s341_ in
- if ((string_startswith _s342_ "medeleg")) then
- match (string_drop _s342_ (projT1 (string_length "medeleg"))) with | s_ => Some (s_) end
- else None.
+Definition get_utvec '(tt : unit) : M (mword 64) :=
+ read_reg utvec_ref >>= fun w__0 : Mtvec => returnm (_get_Mtvec_bits w__0).
-Definition _s336_ (_s337_ : string)
-: option string :=
-
- let _s338_ := _s337_ in
- if ((string_startswith _s338_ "misa")) then
- match (string_drop _s338_ (projT1 (string_length "misa"))) with | s_ => Some (s_) end
- else None.
+Definition set_mtvec (value : mword 64) : M (mword 64) :=
+ read_reg mtvec_ref >>= fun w__0 : Mtvec =>
+ write_reg mtvec_ref (legalize_tvec w__0 value) >>
+ read_reg mtvec_ref >>= fun w__1 : Mtvec => returnm (_get_Mtvec_bits w__1).
-Definition _s332_ (_s333_ : string)
-: option string :=
-
- let _s334_ := _s333_ in
- if ((string_startswith _s334_ "mstatus")) then
- match (string_drop _s334_ (projT1 (string_length "mstatus"))) with | s_ => Some (s_) end
- else None.
+Definition set_stvec (value : mword 64) : M (mword 64) :=
+ read_reg stvec_ref >>= fun w__0 : Mtvec =>
+ write_reg stvec_ref (legalize_tvec w__0 value) >>
+ read_reg stvec_ref >>= fun w__1 : Mtvec => returnm (_get_Mtvec_bits w__1).
-Definition _s328_ (_s329_ : string)
-: option string :=
-
- let _s330_ := _s329_ in
- if ((string_startswith _s330_ "mhartid")) then
- match (string_drop _s330_ (projT1 (string_length "mhartid"))) with | s_ => Some (s_) end
- else None.
+Definition set_utvec (value : mword 64) : M (mword 64) :=
+ read_reg utvec_ref >>= fun w__0 : Mtvec =>
+ write_reg utvec_ref (legalize_tvec w__0 value) >>
+ read_reg utvec_ref >>= fun w__1 : Mtvec => returnm (_get_Mtvec_bits w__1).
+
+Definition update_softfloat_fflags (flags : mword 5) : M (unit) :=
+ write_reg float_fflags_ref (zero_extend flags 64) : M (unit).
+
+Axiom extern_f32Add : forall (_ : mword 3) (_ : mword 32) (_ : mword 32) , unit.
+
+Definition riscv_f32Add (rm : mword 3) (v1 : mword 32) (v2 : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32Add rm v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f32Sub : forall (_ : mword 3) (_ : mword 32) (_ : mword 32) , unit.
+
+Definition riscv_f32Sub (rm : mword 3) (v1 : mword 32) (v2 : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32Sub rm v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f32Mul : forall (_ : mword 3) (_ : mword 32) (_ : mword 32) , unit.
+
+Definition riscv_f32Mul (rm : mword 3) (v1 : mword 32) (v2 : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32Mul rm v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f32Div : forall (_ : mword 3) (_ : mword 32) (_ : mword 32) , unit.
+
+Definition riscv_f32Div (rm : mword 3) (v1 : mword 32) (v2 : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32Div rm v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f64Add : forall (_ : mword 3) (_ : mword 64) (_ : mword 64) , unit.
+
+Definition riscv_f64Add (rm : mword 3) (v1 : mword 64) (v2 : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64Add rm v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f64Sub : forall (_ : mword 3) (_ : mword 64) (_ : mword 64) , unit.
+
+Definition riscv_f64Sub (rm : mword 3) (v1 : mword 64) (v2 : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64Sub rm v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f64Mul : forall (_ : mword 3) (_ : mword 64) (_ : mword 64) , unit.
+
+Definition riscv_f64Mul (rm : mword 3) (v1 : mword 64) (v2 : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64Mul rm v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f64Div : forall (_ : mword 3) (_ : mword 64) (_ : mword 64) , unit.
+
+Definition riscv_f64Div (rm : mword 3) (v1 : mword 64) (v2 : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64Div rm v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f32MulAdd : forall (_ : mword 3) (_ : mword 32) (_ : mword 32) (_ : mword 32) , unit.
+
+Definition riscv_f32MulAdd (rm : mword 3) (v1 : mword 32) (v2 : mword 32) (v3 : mword 32)
+: M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32MulAdd rm v1 v2 v3) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f64MulAdd : forall (_ : mword 3) (_ : mword 64) (_ : mword 64) (_ : mword 64) , unit.
+
+Definition riscv_f64MulAdd (rm : mword 3) (v1 : mword 64) (v2 : mword 64) (v3 : mword 64)
+: M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64MulAdd rm v1 v2 v3) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f32Sqrt : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_f32Sqrt (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32Sqrt rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f64Sqrt : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_f64Sqrt (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64Sqrt rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f32ToI32 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_f32ToI32 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32ToI32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f32ToUi32 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_f32ToUi32 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32ToUi32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_i32ToF32 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_i32ToF32 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_i32ToF32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_ui32ToF32 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_ui32ToF32 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_ui32ToF32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f32ToI64 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_f32ToI64 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f32ToI64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f32ToUi64 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_f32ToUi64 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f32ToUi64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_i64ToF32 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_i64ToF32 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_i64ToF32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_ui64ToF32 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_ui64ToF32 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_ui64ToF32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f64ToI32 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_f64ToI32 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f64ToI32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f64ToUi32 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_f64ToUi32 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f64ToUi32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_i32ToF64 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_i32ToF64 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_i32ToF64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_ui32ToF64 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_ui32ToF64 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_ui32ToF64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f64ToI64 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_f64ToI64 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64ToI64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f64ToUi64 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_f64ToUi64 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64ToUi64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_i64ToF64 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_i64ToF64 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_i64ToF64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_ui64ToF64 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_ui64ToF64 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_ui64ToF64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f32ToF64 : forall (_ : mword 3) (_ : mword 32) , unit.
+
+Definition riscv_f32ToF64 (rm : mword 3) (v : mword 32) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f32ToF64 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f64ToF32 : forall (_ : mword 3) (_ : mword 64) , unit.
+
+Definition riscv_f64ToF32 (rm : mword 3) (v : mword 64) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f64ToF32 rm v) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f32Lt : forall (_ : mword 32) (_ : mword 32) , unit.
+
+Definition riscv_f32Lt (v1 : mword 32) (v2 : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32Lt v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f32Le : forall (_ : mword 32) (_ : mword 32) , unit.
+
+Definition riscv_f32Le (v1 : mword 32) (v2 : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32Le v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f32Eq : forall (_ : mword 32) (_ : mword 32) , unit.
-Definition _s324_ (_s325_ : string)
-: option string :=
-
+Definition riscv_f32Eq (v1 : mword 32) (v2 : mword 32) : M ((mword 5 * mword 32)) :=
+ let '_ := (extern_f32Eq v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, subrange_vec_dec w__1 31 0).
+
+Axiom extern_f64Lt : forall (_ : mword 64) (_ : mword 64) , unit.
+
+Definition riscv_f64Lt (v1 : mword 64) (v2 : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64Lt v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f64Le : forall (_ : mword 64) (_ : mword 64) , unit.
+
+Definition riscv_f64Le (v1 : mword 64) (v2 : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64Le v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Axiom extern_f64Eq : forall (_ : mword 64) (_ : mword 64) , unit.
+
+Definition riscv_f64Eq (v1 : mword 64) (v2 : mword 64) : M ((mword 5 * mword 64)) :=
+ let '_ := (extern_f64Eq v1 v2) : unit in
+ ((read_reg float_fflags_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
+ ((read_reg float_result_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ returnm (subrange_vec_dec w__0 4 0, w__1).
+
+Definition dirty_fd_context '(tt : unit) : M (unit) :=
+ (_set_Mstatus_FS mstatus_ref (extStatus_to_bits Dirty)) >>
+ (_set_Mstatus_SD mstatus_ref ('b"1" : mword 1))
+ : M (unit).
+
+Definition rF (r : Z) `{ArithFact ((0 <=? r) && (r <? 32))} : M (mword 64) :=
+ let l__42 := r in
+ (if sumbool_of_bool (Z.eqb l__42 0) then ((read_reg f0_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 1) then ((read_reg f1_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 2) then ((read_reg f2_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 3) then ((read_reg f3_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 4) then ((read_reg f4_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 5) then ((read_reg f5_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 6) then ((read_reg f6_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 7) then ((read_reg f7_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 8) then ((read_reg f8_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 9) then ((read_reg f9_ref) : M (mword 64)) : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 10) then
+ ((read_reg f10_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 11) then
+ ((read_reg f11_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 12) then
+ ((read_reg f12_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 13) then
+ ((read_reg f13_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 14) then
+ ((read_reg f14_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 15) then
+ ((read_reg f15_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 16) then
+ ((read_reg f16_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 17) then
+ ((read_reg f17_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 18) then
+ ((read_reg f18_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 19) then
+ ((read_reg f19_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 20) then
+ ((read_reg f20_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 21) then
+ ((read_reg f21_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 22) then
+ ((read_reg f22_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 23) then
+ ((read_reg f23_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 24) then
+ ((read_reg f24_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 25) then
+ ((read_reg f25_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 26) then
+ ((read_reg f26_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 27) then
+ ((read_reg f27_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 28) then
+ ((read_reg f28_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 29) then
+ ((read_reg f29_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 30) then
+ ((read_reg f30_ref) : M (mword 64))
+ : M (mword 64)
+ else if sumbool_of_bool (Z.eqb l__42 31) then
+ ((read_reg f31_ref) : M (mword 64))
+ : M (mword 64)
+ else assert_exp' false "invalid floating point register number" >>= fun _ => exit tt) >>= fun v : fregtype =>
+ returnm (fregval_from_freg v).
+
+Definition wF (r : Z) (in_v : mword 64) `{ArithFact ((0 <=? r) && (r <? 32))} : M (unit) :=
+ let v := fregval_into_freg in_v in
+ let l__10 := r in
+ (if sumbool_of_bool (Z.eqb l__10 0) then write_reg f0_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 1) then write_reg f1_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 2) then write_reg f2_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 3) then write_reg f3_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 4) then write_reg f4_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 5) then write_reg f5_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 6) then write_reg f6_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 7) then write_reg f7_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 8) then write_reg f8_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 9) then write_reg f9_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 10) then write_reg f10_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 11) then write_reg f11_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 12) then write_reg f12_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 13) then write_reg f13_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 14) then write_reg f14_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 15) then write_reg f15_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 16) then write_reg f16_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 17) then write_reg f17_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 18) then write_reg f18_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 19) then write_reg f19_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 20) then write_reg f20_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 21) then write_reg f21_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 22) then write_reg f22_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 23) then write_reg f23_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 24) then write_reg f24_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 25) then write_reg f25_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 26) then write_reg f26_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 27) then write_reg f27_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 28) then write_reg f28_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 29) then write_reg f29_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 30) then write_reg f30_ref v : M (unit)
+ else if sumbool_of_bool (Z.eqb l__10 31) then write_reg f31_ref v : M (unit)
+ else assert_exp' false "invalid floating point register number" >>= fun _ => exit tt) >>
+ (dirty_fd_context tt) >>
+ returnm (if get_config_print_reg tt then
+ print_endline
+ (String.append "f"
+ (String.append (string_of_int r) (String.append " <- " (FRegStr v))))
+ else tt).
+
+Definition rF_bits (i : mword 5) : M (mword 64) := (rF (projT1 (uint i))) : M (mword 64).
+
+Definition wF_bits (i : mword 5) (data : mword 64) : M (unit) :=
+ (wF (projT1 (uint i)) data) : M (unit).
+
+Definition freg_name_abi_forwards (arg_ : mword 5) : M (string) :=
+ let b__0 := arg_ in
+ (if eq_vec b__0 ('b"00000" : mword 5) then returnm "ft0"
+ else if eq_vec b__0 ('b"00001" : mword 5) then returnm "ft1"
+ else if eq_vec b__0 ('b"00010" : mword 5) then returnm "ft2"
+ else if eq_vec b__0 ('b"00011" : mword 5) then returnm "ft3"
+ else if eq_vec b__0 ('b"00100" : mword 5) then returnm "ft4"
+ else if eq_vec b__0 ('b"00101" : mword 5) then returnm "ft5"
+ else if eq_vec b__0 ('b"00110" : mword 5) then returnm "ft6"
+ else if eq_vec b__0 ('b"00111" : mword 5) then returnm "ft7"
+ else if eq_vec b__0 ('b"01000" : mword 5) then returnm "fs0"
+ else if eq_vec b__0 ('b"01001" : mword 5) then returnm "fs1"
+ else if eq_vec b__0 ('b"01010" : mword 5) then returnm "fa0"
+ else if eq_vec b__0 ('b"01011" : mword 5) then returnm "fa1"
+ else if eq_vec b__0 ('b"01100" : mword 5) then returnm "fa2"
+ else if eq_vec b__0 ('b"01101" : mword 5) then returnm "fa3"
+ else if eq_vec b__0 ('b"01110" : mword 5) then returnm "fa4"
+ else if eq_vec b__0 ('b"01111" : mword 5) then returnm "fa5"
+ else if eq_vec b__0 ('b"10000" : mword 5) then returnm "fa6"
+ else if eq_vec b__0 ('b"10001" : mword 5) then returnm "fa7"
+ else if eq_vec b__0 ('b"10010" : mword 5) then returnm "fs2"
+ else if eq_vec b__0 ('b"10011" : mword 5) then returnm "fs3"
+ else if eq_vec b__0 ('b"10100" : mword 5) then returnm "fs4"
+ else if eq_vec b__0 ('b"10101" : mword 5) then returnm "fs5"
+ else if eq_vec b__0 ('b"10110" : mword 5) then returnm "fs6"
+ else if eq_vec b__0 ('b"10111" : mword 5) then returnm "fs7"
+ else if eq_vec b__0 ('b"11000" : mword 5) then returnm "fs8"
+ else if eq_vec b__0 ('b"11001" : mword 5) then returnm "fs9"
+ else if eq_vec b__0 ('b"11010" : mword 5) then returnm "fs10"
+ else if eq_vec b__0 ('b"11011" : mword 5) then returnm "fs11"
+ else if eq_vec b__0 ('b"11100" : mword 5) then returnm "ft8"
+ else if eq_vec b__0 ('b"11101" : mword 5) then returnm "ft9"
+ else if eq_vec b__0 ('b"11110" : mword 5) then returnm "ft10"
+ else if eq_vec b__0 ('b"11111" : mword 5) then returnm "ft11"
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string).
+
+Definition freg_name_abi_backwards (arg_ : string) : M (mword 5) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "ft0" then returnm ('b"00000" : mword 5)
+ else if generic_eq p0_ "ft1" then returnm ('b"00001" : mword 5)
+ else if generic_eq p0_ "ft2" then returnm ('b"00010" : mword 5)
+ else if generic_eq p0_ "ft3" then returnm ('b"00011" : mword 5)
+ else if generic_eq p0_ "ft4" then returnm ('b"00100" : mword 5)
+ else if generic_eq p0_ "ft5" then returnm ('b"00101" : mword 5)
+ else if generic_eq p0_ "ft6" then returnm ('b"00110" : mword 5)
+ else if generic_eq p0_ "ft7" then returnm ('b"00111" : mword 5)
+ else if generic_eq p0_ "fs0" then returnm ('b"01000" : mword 5)
+ else if generic_eq p0_ "fs1" then returnm ('b"01001" : mword 5)
+ else if generic_eq p0_ "fa0" then returnm ('b"01010" : mword 5)
+ else if generic_eq p0_ "fa1" then returnm ('b"01011" : mword 5)
+ else if generic_eq p0_ "fa2" then returnm ('b"01100" : mword 5)
+ else if generic_eq p0_ "fa3" then returnm ('b"01101" : mword 5)
+ else if generic_eq p0_ "fa4" then returnm ('b"01110" : mword 5)
+ else if generic_eq p0_ "fa5" then returnm ('b"01111" : mword 5)
+ else if generic_eq p0_ "fa6" then returnm ('b"10000" : mword 5)
+ else if generic_eq p0_ "fa7" then returnm ('b"10001" : mword 5)
+ else if generic_eq p0_ "fs2" then returnm ('b"10010" : mword 5)
+ else if generic_eq p0_ "fs3" then returnm ('b"10011" : mword 5)
+ else if generic_eq p0_ "fs4" then returnm ('b"10100" : mword 5)
+ else if generic_eq p0_ "fs5" then returnm ('b"10101" : mword 5)
+ else if generic_eq p0_ "fs6" then returnm ('b"10110" : mword 5)
+ else if generic_eq p0_ "fs7" then returnm ('b"10111" : mword 5)
+ else if generic_eq p0_ "fs8" then returnm ('b"11000" : mword 5)
+ else if generic_eq p0_ "fs9" then returnm ('b"11001" : mword 5)
+ else if generic_eq p0_ "fs10" then returnm ('b"11010" : mword 5)
+ else if generic_eq p0_ "fs11" then returnm ('b"11011" : mword 5)
+ else if generic_eq p0_ "ft8" then returnm ('b"11100" : mword 5)
+ else if generic_eq p0_ "ft9" then returnm ('b"11101" : mword 5)
+ else if generic_eq p0_ "ft10" then returnm ('b"11110" : mword 5)
+ else if generic_eq p0_ "ft11" then returnm ('b"11111" : mword 5)
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 5).
+
+Definition freg_name_abi_forwards_matches (arg_ : mword 5) : bool :=
+ let b__0 := arg_ in
+ if eq_vec b__0 ('b"00000" : mword 5) then true
+ else if eq_vec b__0 ('b"00001" : mword 5) then true
+ else if eq_vec b__0 ('b"00010" : mword 5) then true
+ else if eq_vec b__0 ('b"00011" : mword 5) then true
+ else if eq_vec b__0 ('b"00100" : mword 5) then true
+ else if eq_vec b__0 ('b"00101" : mword 5) then true
+ else if eq_vec b__0 ('b"00110" : mword 5) then true
+ else if eq_vec b__0 ('b"00111" : mword 5) then true
+ else if eq_vec b__0 ('b"01000" : mword 5) then true
+ else if eq_vec b__0 ('b"01001" : mword 5) then true
+ else if eq_vec b__0 ('b"01010" : mword 5) then true
+ else if eq_vec b__0 ('b"01011" : mword 5) then true
+ else if eq_vec b__0 ('b"01100" : mword 5) then true
+ else if eq_vec b__0 ('b"01101" : mword 5) then true
+ else if eq_vec b__0 ('b"01110" : mword 5) then true
+ else if eq_vec b__0 ('b"01111" : mword 5) then true
+ else if eq_vec b__0 ('b"10000" : mword 5) then true
+ else if eq_vec b__0 ('b"10001" : mword 5) then true
+ else if eq_vec b__0 ('b"10010" : mword 5) then true
+ else if eq_vec b__0 ('b"10011" : mword 5) then true
+ else if eq_vec b__0 ('b"10100" : mword 5) then true
+ else if eq_vec b__0 ('b"10101" : mword 5) then true
+ else if eq_vec b__0 ('b"10110" : mword 5) then true
+ else if eq_vec b__0 ('b"10111" : mword 5) then true
+ else if eq_vec b__0 ('b"11000" : mword 5) then true
+ else if eq_vec b__0 ('b"11001" : mword 5) then true
+ else if eq_vec b__0 ('b"11010" : mword 5) then true
+ else if eq_vec b__0 ('b"11011" : mword 5) then true
+ else if eq_vec b__0 ('b"11100" : mword 5) then true
+ else if eq_vec b__0 ('b"11101" : mword 5) then true
+ else if eq_vec b__0 ('b"11110" : mword 5) then true
+ else if eq_vec b__0 ('b"11111" : mword 5) then true
+ else false.
+
+Definition freg_name_abi_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "ft0" then true
+ else if generic_eq p0_ "ft1" then true
+ else if generic_eq p0_ "ft2" then true
+ else if generic_eq p0_ "ft3" then true
+ else if generic_eq p0_ "ft4" then true
+ else if generic_eq p0_ "ft5" then true
+ else if generic_eq p0_ "ft6" then true
+ else if generic_eq p0_ "ft7" then true
+ else if generic_eq p0_ "fs0" then true
+ else if generic_eq p0_ "fs1" then true
+ else if generic_eq p0_ "fa0" then true
+ else if generic_eq p0_ "fa1" then true
+ else if generic_eq p0_ "fa2" then true
+ else if generic_eq p0_ "fa3" then true
+ else if generic_eq p0_ "fa4" then true
+ else if generic_eq p0_ "fa5" then true
+ else if generic_eq p0_ "fa6" then true
+ else if generic_eq p0_ "fa7" then true
+ else if generic_eq p0_ "fs2" then true
+ else if generic_eq p0_ "fs3" then true
+ else if generic_eq p0_ "fs4" then true
+ else if generic_eq p0_ "fs5" then true
+ else if generic_eq p0_ "fs6" then true
+ else if generic_eq p0_ "fs7" then true
+ else if generic_eq p0_ "fs8" then true
+ else if generic_eq p0_ "fs9" then true
+ else if generic_eq p0_ "fs10" then true
+ else if generic_eq p0_ "fs11" then true
+ else if generic_eq p0_ "ft8" then true
+ else if generic_eq p0_ "ft9" then true
+ else if generic_eq p0_ "ft10" then true
+ else if generic_eq p0_ "ft11" then true
+ else false.
+
+Definition _s324_ (_s325_ : string) : option string :=
let _s326_ := _s325_ in
- if ((string_startswith _s326_ "mimpid")) then
- match (string_drop _s326_ (projT1 (string_length "mimpid"))) with | s_ => Some (s_) end
+ if string_startswith _s326_ "ft11" then
+ match (string_drop _s326_ (projT1 (string_length "ft11"))) with | s_ => Some s_ end
else None.
-Definition _s320_ (_s321_ : string)
-: option string :=
-
+Definition _s320_ (_s321_ : string) : option string :=
let _s322_ := _s321_ in
- if ((string_startswith _s322_ "marchid")) then
- match (string_drop _s322_ (projT1 (string_length "marchid"))) with | s_ => Some (s_) end
+ if string_startswith _s322_ "ft10" then
+ match (string_drop _s322_ (projT1 (string_length "ft10"))) with | s_ => Some s_ end
else None.
-Definition _s316_ (_s317_ : string)
-: option string :=
-
+Definition _s316_ (_s317_ : string) : option string :=
let _s318_ := _s317_ in
- if ((string_startswith _s318_ "mvendorid")) then
- match (string_drop _s318_ (projT1 (string_length "mvendorid"))) with | s_ => Some (s_) end
+ if string_startswith _s318_ "ft9" then
+ match (string_drop _s318_ (projT1 (string_length "ft9"))) with | s_ => Some s_ end
else None.
-Definition _s312_ (_s313_ : string)
-: option string :=
-
+Definition _s312_ (_s313_ : string) : option string :=
let _s314_ := _s313_ in
- if ((string_startswith _s314_ "satp")) then
- match (string_drop _s314_ (projT1 (string_length "satp"))) with | s_ => Some (s_) end
+ if string_startswith _s314_ "ft8" then
+ match (string_drop _s314_ (projT1 (string_length "ft8"))) with | s_ => Some s_ end
else None.
-Definition _s308_ (_s309_ : string)
-: option string :=
-
+Definition _s308_ (_s309_ : string) : option string :=
let _s310_ := _s309_ in
- if ((string_startswith _s310_ "sip")) then
- match (string_drop _s310_ (projT1 (string_length "sip"))) with | s_ => Some (s_) end
+ if string_startswith _s310_ "fs11" then
+ match (string_drop _s310_ (projT1 (string_length "fs11"))) with | s_ => Some s_ end
else None.
-Definition _s304_ (_s305_ : string)
-: option string :=
-
+Definition _s304_ (_s305_ : string) : option string :=
let _s306_ := _s305_ in
- if ((string_startswith _s306_ "stval")) then
- match (string_drop _s306_ (projT1 (string_length "stval"))) with | s_ => Some (s_) end
+ if string_startswith _s306_ "fs10" then
+ match (string_drop _s306_ (projT1 (string_length "fs10"))) with | s_ => Some s_ end
else None.
-Definition _s300_ (_s301_ : string)
-: option string :=
-
+Definition _s300_ (_s301_ : string) : option string :=
let _s302_ := _s301_ in
- if ((string_startswith _s302_ "scause")) then
- match (string_drop _s302_ (projT1 (string_length "scause"))) with | s_ => Some (s_) end
+ if string_startswith _s302_ "fs9" then
+ match (string_drop _s302_ (projT1 (string_length "fs9"))) with | s_ => Some s_ end
else None.
-Definition _s296_ (_s297_ : string)
-: option string :=
-
+Definition _s296_ (_s297_ : string) : option string :=
let _s298_ := _s297_ in
- if ((string_startswith _s298_ "sepc")) then
- match (string_drop _s298_ (projT1 (string_length "sepc"))) with | s_ => Some (s_) end
+ if string_startswith _s298_ "fs8" then
+ match (string_drop _s298_ (projT1 (string_length "fs8"))) with | s_ => Some s_ end
else None.
-Definition _s292_ (_s293_ : string)
-: option string :=
-
+Definition _s292_ (_s293_ : string) : option string :=
let _s294_ := _s293_ in
- if ((string_startswith _s294_ "sscratch")) then
- match (string_drop _s294_ (projT1 (string_length "sscratch"))) with | s_ => Some (s_) end
+ if string_startswith _s294_ "fs7" then
+ match (string_drop _s294_ (projT1 (string_length "fs7"))) with | s_ => Some s_ end
else None.
-Definition _s288_ (_s289_ : string)
-: option string :=
-
+Definition _s288_ (_s289_ : string) : option string :=
let _s290_ := _s289_ in
- if ((string_startswith _s290_ "scounteren")) then
- match (string_drop _s290_ (projT1 (string_length "scounteren"))) with | s_ => Some (s_) end
+ if string_startswith _s290_ "fs6" then
+ match (string_drop _s290_ (projT1 (string_length "fs6"))) with | s_ => Some s_ end
else None.
-Definition _s284_ (_s285_ : string)
-: option string :=
-
+Definition _s284_ (_s285_ : string) : option string :=
let _s286_ := _s285_ in
- if ((string_startswith _s286_ "stvec")) then
- match (string_drop _s286_ (projT1 (string_length "stvec"))) with | s_ => Some (s_) end
+ if string_startswith _s286_ "fs5" then
+ match (string_drop _s286_ (projT1 (string_length "fs5"))) with | s_ => Some s_ end
else None.
-Definition _s280_ (_s281_ : string)
-: option string :=
-
+Definition _s280_ (_s281_ : string) : option string :=
let _s282_ := _s281_ in
- if ((string_startswith _s282_ "sie")) then
- match (string_drop _s282_ (projT1 (string_length "sie"))) with | s_ => Some (s_) end
+ if string_startswith _s282_ "fs4" then
+ match (string_drop _s282_ (projT1 (string_length "fs4"))) with | s_ => Some s_ end
else None.
-Definition _s276_ (_s277_ : string)
-: option string :=
-
+Definition _s276_ (_s277_ : string) : option string :=
let _s278_ := _s277_ in
- if ((string_startswith _s278_ "sideleg")) then
- match (string_drop _s278_ (projT1 (string_length "sideleg"))) with | s_ => Some (s_) end
+ if string_startswith _s278_ "fs3" then
+ match (string_drop _s278_ (projT1 (string_length "fs3"))) with | s_ => Some s_ end
else None.
-Definition _s272_ (_s273_ : string)
-: option string :=
-
+Definition _s272_ (_s273_ : string) : option string :=
let _s274_ := _s273_ in
- if ((string_startswith _s274_ "sedeleg")) then
- match (string_drop _s274_ (projT1 (string_length "sedeleg"))) with | s_ => Some (s_) end
+ if string_startswith _s274_ "fs2" then
+ match (string_drop _s274_ (projT1 (string_length "fs2"))) with | s_ => Some s_ end
else None.
-Definition _s268_ (_s269_ : string)
-: option string :=
-
+Definition _s268_ (_s269_ : string) : option string :=
let _s270_ := _s269_ in
- if ((string_startswith _s270_ "sstatus")) then
- match (string_drop _s270_ (projT1 (string_length "sstatus"))) with | s_ => Some (s_) end
+ if string_startswith _s270_ "fa7" then
+ match (string_drop _s270_ (projT1 (string_length "fa7"))) with | s_ => Some s_ end
else None.
-Definition _s264_ (_s265_ : string)
-: option string :=
-
+Definition _s264_ (_s265_ : string) : option string :=
let _s266_ := _s265_ in
- if ((string_startswith _s266_ "instreth")) then
- match (string_drop _s266_ (projT1 (string_length "instreth"))) with | s_ => Some (s_) end
+ if string_startswith _s266_ "fa6" then
+ match (string_drop _s266_ (projT1 (string_length "fa6"))) with | s_ => Some s_ end
else None.
-Definition _s260_ (_s261_ : string)
-: option string :=
-
+Definition _s260_ (_s261_ : string) : option string :=
let _s262_ := _s261_ in
- if ((string_startswith _s262_ "timeh")) then
- match (string_drop _s262_ (projT1 (string_length "timeh"))) with | s_ => Some (s_) end
+ if string_startswith _s262_ "fa5" then
+ match (string_drop _s262_ (projT1 (string_length "fa5"))) with | s_ => Some s_ end
else None.
-Definition _s256_ (_s257_ : string)
-: option string :=
-
+Definition _s256_ (_s257_ : string) : option string :=
let _s258_ := _s257_ in
- if ((string_startswith _s258_ "cycleh")) then
- match (string_drop _s258_ (projT1 (string_length "cycleh"))) with | s_ => Some (s_) end
+ if string_startswith _s258_ "fa4" then
+ match (string_drop _s258_ (projT1 (string_length "fa4"))) with | s_ => Some s_ end
else None.
-Definition _s252_ (_s253_ : string)
-: option string :=
-
+Definition _s252_ (_s253_ : string) : option string :=
let _s254_ := _s253_ in
- if ((string_startswith _s254_ "instret")) then
- match (string_drop _s254_ (projT1 (string_length "instret"))) with | s_ => Some (s_) end
+ if string_startswith _s254_ "fa3" then
+ match (string_drop _s254_ (projT1 (string_length "fa3"))) with | s_ => Some s_ end
else None.
-Definition _s248_ (_s249_ : string)
-: option string :=
-
+Definition _s248_ (_s249_ : string) : option string :=
let _s250_ := _s249_ in
- if ((string_startswith _s250_ "time")) then
- match (string_drop _s250_ (projT1 (string_length "time"))) with | s_ => Some (s_) end
+ if string_startswith _s250_ "fa2" then
+ match (string_drop _s250_ (projT1 (string_length "fa2"))) with | s_ => Some s_ end
else None.
-Definition _s244_ (_s245_ : string)
-: option string :=
-
+Definition _s244_ (_s245_ : string) : option string :=
let _s246_ := _s245_ in
- if ((string_startswith _s246_ "cycle")) then
- match (string_drop _s246_ (projT1 (string_length "cycle"))) with | s_ => Some (s_) end
+ if string_startswith _s246_ "fa1" then
+ match (string_drop _s246_ (projT1 (string_length "fa1"))) with | s_ => Some s_ end
else None.
-Definition _s240_ (_s241_ : string)
-: option string :=
-
+Definition _s240_ (_s241_ : string) : option string :=
let _s242_ := _s241_ in
- if ((string_startswith _s242_ "fcsr")) then
- match (string_drop _s242_ (projT1 (string_length "fcsr"))) with | s_ => Some (s_) end
+ if string_startswith _s242_ "fa0" then
+ match (string_drop _s242_ (projT1 (string_length "fa0"))) with | s_ => Some s_ end
else None.
-Definition _s236_ (_s237_ : string)
-: option string :=
-
+Definition _s236_ (_s237_ : string) : option string :=
let _s238_ := _s237_ in
- if ((string_startswith _s238_ "frm")) then
- match (string_drop _s238_ (projT1 (string_length "frm"))) with | s_ => Some (s_) end
+ if string_startswith _s238_ "fs1" then
+ match (string_drop _s238_ (projT1 (string_length "fs1"))) with | s_ => Some s_ end
else None.
-Definition _s232_ (_s233_ : string)
-: option string :=
-
+Definition _s232_ (_s233_ : string) : option string :=
let _s234_ := _s233_ in
- if ((string_startswith _s234_ "fflags")) then
- match (string_drop _s234_ (projT1 (string_length "fflags"))) with | s_ => Some (s_) end
+ if string_startswith _s234_ "fs0" then
+ match (string_drop _s234_ (projT1 (string_length "fs0"))) with | s_ => Some s_ end
else None.
-Definition _s228_ (_s229_ : string)
-: option string :=
-
+Definition _s228_ (_s229_ : string) : option string :=
let _s230_ := _s229_ in
- if ((string_startswith _s230_ "uip")) then
- match (string_drop _s230_ (projT1 (string_length "uip"))) with | s_ => Some (s_) end
+ if string_startswith _s230_ "ft7" then
+ match (string_drop _s230_ (projT1 (string_length "ft7"))) with | s_ => Some s_ end
else None.
-Definition _s224_ (_s225_ : string)
-: option string :=
-
+Definition _s224_ (_s225_ : string) : option string :=
let _s226_ := _s225_ in
- if ((string_startswith _s226_ "utval")) then
- match (string_drop _s226_ (projT1 (string_length "utval"))) with | s_ => Some (s_) end
+ if string_startswith _s226_ "ft6" then
+ match (string_drop _s226_ (projT1 (string_length "ft6"))) with | s_ => Some s_ end
else None.
-Definition _s220_ (_s221_ : string)
-: option string :=
-
+Definition _s220_ (_s221_ : string) : option string :=
let _s222_ := _s221_ in
- if ((string_startswith _s222_ "ucause")) then
- match (string_drop _s222_ (projT1 (string_length "ucause"))) with | s_ => Some (s_) end
+ if string_startswith _s222_ "ft5" then
+ match (string_drop _s222_ (projT1 (string_length "ft5"))) with | s_ => Some s_ end
else None.
-Definition _s216_ (_s217_ : string)
-: option string :=
-
+Definition _s216_ (_s217_ : string) : option string :=
let _s218_ := _s217_ in
- if ((string_startswith _s218_ "uepc")) then
- match (string_drop _s218_ (projT1 (string_length "uepc"))) with | s_ => Some (s_) end
+ if string_startswith _s218_ "ft4" then
+ match (string_drop _s218_ (projT1 (string_length "ft4"))) with | s_ => Some s_ end
else None.
-Definition _s212_ (_s213_ : string)
-: option string :=
-
+Definition _s212_ (_s213_ : string) : option string :=
let _s214_ := _s213_ in
- if ((string_startswith _s214_ "uscratch")) then
- match (string_drop _s214_ (projT1 (string_length "uscratch"))) with | s_ => Some (s_) end
+ if string_startswith _s214_ "ft3" then
+ match (string_drop _s214_ (projT1 (string_length "ft3"))) with | s_ => Some s_ end
else None.
-Definition _s208_ (_s209_ : string)
-: option string :=
-
+Definition _s208_ (_s209_ : string) : option string :=
let _s210_ := _s209_ in
- if ((string_startswith _s210_ "utvec")) then
- match (string_drop _s210_ (projT1 (string_length "utvec"))) with | s_ => Some (s_) end
+ if string_startswith _s210_ "ft2" then
+ match (string_drop _s210_ (projT1 (string_length "ft2"))) with | s_ => Some s_ end
else None.
-Definition _s204_ (_s205_ : string)
-: option string :=
-
+Definition _s204_ (_s205_ : string) : option string :=
let _s206_ := _s205_ in
- if ((string_startswith _s206_ "uie")) then
- match (string_drop _s206_ (projT1 (string_length "uie"))) with | s_ => Some (s_) end
+ if string_startswith _s206_ "ft1" then
+ match (string_drop _s206_ (projT1 (string_length "ft1"))) with | s_ => Some s_ end
else None.
-Definition _s200_ (_s201_ : string)
-: option string :=
-
+Definition _s200_ (_s201_ : string) : option string :=
let _s202_ := _s201_ in
- if ((string_startswith _s202_ "ustatus")) then
- match (string_drop _s202_ (projT1 (string_length "ustatus"))) with | s_ => Some (s_) end
+ if string_startswith _s202_ "ft0" then
+ match (string_drop _s202_ (projT1 (string_length "ft0"))) with | s_ => Some s_ end
else None.
-Definition csr_name_map_matches_prefix (arg_ : string)
-: M (option ((mword 12 * {n : Z & ArithFact (n >= 0)}))) :=
-
+Definition freg_name_abi_matches_prefix (arg_ : string)
+: M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)}))) :=
let _s203_ := arg_ in
- (if ((match (_s200_ _s203_) with | Some (s_) => true | _ => false end)) then
+ (if match (_s200_ _s203_) with | Some s_ => true | _ => false end then
(match (_s200_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s204_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s204_ _s203_) with | Some s_ => true | _ => false end then
(match (_s204_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s208_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s208_ _s203_) with | Some s_ => true | _ => false end then
(match (_s208_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s212_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s212_ _s203_) with | Some s_ => true | _ => false end then
(match (_s212_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s216_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s216_ _s203_) with | Some s_ => true | _ => false end then
(match (_s216_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s220_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s220_ _s203_) with | Some s_ => true | _ => false end then
(match (_s220_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s224_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s224_ _s203_) with | Some s_ => true | _ => false end then
(match (_s224_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s228_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s228_ _s203_) with | Some s_ => true | _ => false end then
(match (_s228_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"00111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s232_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s232_ _s203_) with | Some s_ => true | _ => false end then
(match (_s232_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s236_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s236_ _s203_) with | Some s_ => true | _ => false end then
(match (_s236_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s240_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s240_ _s203_) with | Some s_ => true | _ => false end then
(match (_s240_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s244_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s244_ _s203_) with | Some s_ => true | _ => false end then
(match (_s244_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s248_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s248_ _s203_) with | Some s_ => true | _ => false end then
(match (_s248_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s252_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s252_ _s203_) with | Some s_ => true | _ => false end then
(match (_s252_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s256_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s256_ _s203_) with | Some s_ => true | _ => false end then
(match (_s256_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s260_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s260_ _s203_) with | Some s_ => true | _ => false end then
(match (_s260_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"01111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s264_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s264_ _s203_) with | Some s_ => true | _ => false end then
(match (_s264_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s268_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s268_ _s203_) with | Some s_ => true | _ => false end then
(match (_s268_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s272_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s272_ _s203_) with | Some s_ => true | _ => false end then
(match (_s272_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s276_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s276_ _s203_) with | Some s_ => true | _ => false end then
(match (_s276_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s280_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s280_ _s203_) with | Some s_ => true | _ => false end then
(match (_s280_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s284_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s284_ _s203_) with | Some s_ => true | _ => false end then
(match (_s284_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s288_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s288_ _s203_) with | Some s_ => true | _ => false end then
(match (_s288_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s292_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s292_ _s203_) with | Some s_ => true | _ => false end then
(match (_s292_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"10111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s296_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s296_ _s203_) with | Some s_ => true | _ => false end then
(match (_s296_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s300_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s300_ _s203_) with | Some s_ => true | _ => false end then
(match (_s300_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s304_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s304_ _s203_) with | Some s_ => true | _ => false end then
(match (_s304_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s308_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s308_ _s203_) with | Some s_ => true | _ => false end then
(match (_s308_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s312_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s312_ _s203_) with | Some s_ => true | _ => false end then
(match (_s312_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s316_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s316_ _s203_) with | Some s_ => true | _ => false end then
(match (_s316_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s320_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s320_ _s203_) with | Some s_ => true | _ => false end then
(match (_s320_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s324_ _s203_) with | Some (s_) => true | _ => false end)) then
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s324_ _s203_) with | Some s_ => true | _ => false end then
(match (_s324_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s328_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s328_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s332_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s332_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s336_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s336_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s340_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s340_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s344_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s344_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s348_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s348_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s352_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s352_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s356_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s356_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s360_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s360_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s364_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s364_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s368_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s368_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s372_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s372_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s376_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s376_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s380_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s380_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s384_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s384_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s388_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s388_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s392_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s392_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s396_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s396_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s400_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s400_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s404_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s404_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s408_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s408_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s412_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s412_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s416_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s416_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s420_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s420_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s424_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s424_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s428_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s428_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s432_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s432_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s436_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s436_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s440_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s440_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s444_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s444_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s448_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s448_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s452_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s452_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s456_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s456_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s460_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s460_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s464_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s464_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s468_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s468_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s472_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s472_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s476_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s476_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s480_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s480_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s484_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s484_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s488_ _s203_) with | Some (s_) => true | _ => false end)) then
- (match (_s488_ _s203_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
+ | Some s_ =>
+ returnm (Some
+ ('b"11111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((mword 12 * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((mword 12 * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)}))).
-Definition Mk_Ustatus (v : mword 64)
-: Ustatus :=
-
- {| Ustatus_Ustatus_chunk_0 := (subrange_vec_dec v 63 0) |}.
+Definition freg_name_forwards (arg_ : mword 5) : M (string) :=
+ let b__0 := arg_ in
+ (if eq_vec b__0 ('b"00000" : mword 5) then returnm "ft0"
+ else if eq_vec b__0 ('b"00001" : mword 5) then returnm "ft1"
+ else if eq_vec b__0 ('b"00010" : mword 5) then returnm "ft2"
+ else if eq_vec b__0 ('b"00011" : mword 5) then returnm "ft3"
+ else if eq_vec b__0 ('b"00100" : mword 5) then returnm "ft4"
+ else if eq_vec b__0 ('b"00101" : mword 5) then returnm "ft5"
+ else if eq_vec b__0 ('b"00110" : mword 5) then returnm "ft6"
+ else if eq_vec b__0 ('b"00111" : mword 5) then returnm "ft7"
+ else if eq_vec b__0 ('b"01000" : mword 5) then returnm "fs0"
+ else if eq_vec b__0 ('b"01001" : mword 5) then returnm "fs1"
+ else if eq_vec b__0 ('b"01010" : mword 5) then returnm "fa0"
+ else if eq_vec b__0 ('b"01011" : mword 5) then returnm "fa1"
+ else if eq_vec b__0 ('b"01100" : mword 5) then returnm "fa2"
+ else if eq_vec b__0 ('b"01101" : mword 5) then returnm "fa3"
+ else if eq_vec b__0 ('b"01110" : mword 5) then returnm "fa4"
+ else if eq_vec b__0 ('b"01111" : mword 5) then returnm "fa5"
+ else if eq_vec b__0 ('b"10000" : mword 5) then returnm "fa6"
+ else if eq_vec b__0 ('b"10001" : mword 5) then returnm "fa7"
+ else if eq_vec b__0 ('b"10010" : mword 5) then returnm "fs2"
+ else if eq_vec b__0 ('b"10011" : mword 5) then returnm "fs3"
+ else if eq_vec b__0 ('b"10100" : mword 5) then returnm "fs4"
+ else if eq_vec b__0 ('b"10101" : mword 5) then returnm "fs5"
+ else if eq_vec b__0 ('b"10110" : mword 5) then returnm "fs6"
+ else if eq_vec b__0 ('b"10111" : mword 5) then returnm "fs7"
+ else if eq_vec b__0 ('b"11000" : mword 5) then returnm "fs8"
+ else if eq_vec b__0 ('b"11001" : mword 5) then returnm "fs9"
+ else if eq_vec b__0 ('b"11010" : mword 5) then returnm "fs10"
+ else if eq_vec b__0 ('b"11011" : mword 5) then returnm "fs11"
+ else if eq_vec b__0 ('b"11100" : mword 5) then returnm "ft8"
+ else if eq_vec b__0 ('b"11101" : mword 5) then returnm "ft9"
+ else if eq_vec b__0 ('b"11110" : mword 5) then returnm "ft10"
+ else if eq_vec b__0 ('b"11111" : mword 5) then returnm "ft11"
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string).
-Definition _get_Ustatus_bits (v : Ustatus)
-: mword 64 :=
-
- subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 63 0.
+Definition freg_name_backwards (arg_ : string) : M (mword 5) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "ft0" then returnm ('b"00000" : mword 5)
+ else if generic_eq p0_ "ft1" then returnm ('b"00001" : mword 5)
+ else if generic_eq p0_ "ft2" then returnm ('b"00010" : mword 5)
+ else if generic_eq p0_ "ft3" then returnm ('b"00011" : mword 5)
+ else if generic_eq p0_ "ft4" then returnm ('b"00100" : mword 5)
+ else if generic_eq p0_ "ft5" then returnm ('b"00101" : mword 5)
+ else if generic_eq p0_ "ft6" then returnm ('b"00110" : mword 5)
+ else if generic_eq p0_ "ft7" then returnm ('b"00111" : mword 5)
+ else if generic_eq p0_ "fs0" then returnm ('b"01000" : mword 5)
+ else if generic_eq p0_ "fs1" then returnm ('b"01001" : mword 5)
+ else if generic_eq p0_ "fa0" then returnm ('b"01010" : mword 5)
+ else if generic_eq p0_ "fa1" then returnm ('b"01011" : mword 5)
+ else if generic_eq p0_ "fa2" then returnm ('b"01100" : mword 5)
+ else if generic_eq p0_ "fa3" then returnm ('b"01101" : mword 5)
+ else if generic_eq p0_ "fa4" then returnm ('b"01110" : mword 5)
+ else if generic_eq p0_ "fa5" then returnm ('b"01111" : mword 5)
+ else if generic_eq p0_ "fa6" then returnm ('b"10000" : mword 5)
+ else if generic_eq p0_ "fa7" then returnm ('b"10001" : mword 5)
+ else if generic_eq p0_ "fs2" then returnm ('b"10010" : mword 5)
+ else if generic_eq p0_ "fs3" then returnm ('b"10011" : mword 5)
+ else if generic_eq p0_ "fs4" then returnm ('b"10100" : mword 5)
+ else if generic_eq p0_ "fs5" then returnm ('b"10101" : mword 5)
+ else if generic_eq p0_ "fs6" then returnm ('b"10110" : mword 5)
+ else if generic_eq p0_ "fs7" then returnm ('b"10111" : mword 5)
+ else if generic_eq p0_ "fs8" then returnm ('b"11000" : mword 5)
+ else if generic_eq p0_ "fs9" then returnm ('b"11001" : mword 5)
+ else if generic_eq p0_ "fs10" then returnm ('b"11010" : mword 5)
+ else if generic_eq p0_ "fs11" then returnm ('b"11011" : mword 5)
+ else if generic_eq p0_ "ft8" then returnm ('b"11100" : mword 5)
+ else if generic_eq p0_ "ft9" then returnm ('b"11101" : mword 5)
+ else if generic_eq p0_ "ft10" then returnm ('b"11110" : mword 5)
+ else if generic_eq p0_ "ft11" then returnm ('b"11111" : mword 5)
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 5).
-Definition _set_Ustatus_bits (r_ref : register_ref regstate register_value Ustatus) (v : mword 64)
-: M (unit) :=
-
- (reg_deref r_ref) >>= fun r =>
- let r :=
- {[ r with
- Ustatus_Ustatus_chunk_0 :=
- (update_subrange_vec_dec r.(Ustatus_Ustatus_chunk_0) 63 0 (subrange_vec_dec v 63 0)) ]}
- : Ustatus in
- write_reg r_ref r
- : M (unit).
+Definition freg_name_forwards_matches (arg_ : mword 5) : bool :=
+ let b__0 := arg_ in
+ if eq_vec b__0 ('b"00000" : mword 5) then true
+ else if eq_vec b__0 ('b"00001" : mword 5) then true
+ else if eq_vec b__0 ('b"00010" : mword 5) then true
+ else if eq_vec b__0 ('b"00011" : mword 5) then true
+ else if eq_vec b__0 ('b"00100" : mword 5) then true
+ else if eq_vec b__0 ('b"00101" : mword 5) then true
+ else if eq_vec b__0 ('b"00110" : mword 5) then true
+ else if eq_vec b__0 ('b"00111" : mword 5) then true
+ else if eq_vec b__0 ('b"01000" : mword 5) then true
+ else if eq_vec b__0 ('b"01001" : mword 5) then true
+ else if eq_vec b__0 ('b"01010" : mword 5) then true
+ else if eq_vec b__0 ('b"01011" : mword 5) then true
+ else if eq_vec b__0 ('b"01100" : mword 5) then true
+ else if eq_vec b__0 ('b"01101" : mword 5) then true
+ else if eq_vec b__0 ('b"01110" : mword 5) then true
+ else if eq_vec b__0 ('b"01111" : mword 5) then true
+ else if eq_vec b__0 ('b"10000" : mword 5) then true
+ else if eq_vec b__0 ('b"10001" : mword 5) then true
+ else if eq_vec b__0 ('b"10010" : mword 5) then true
+ else if eq_vec b__0 ('b"10011" : mword 5) then true
+ else if eq_vec b__0 ('b"10100" : mword 5) then true
+ else if eq_vec b__0 ('b"10101" : mword 5) then true
+ else if eq_vec b__0 ('b"10110" : mword 5) then true
+ else if eq_vec b__0 ('b"10111" : mword 5) then true
+ else if eq_vec b__0 ('b"11000" : mword 5) then true
+ else if eq_vec b__0 ('b"11001" : mword 5) then true
+ else if eq_vec b__0 ('b"11010" : mword 5) then true
+ else if eq_vec b__0 ('b"11011" : mword 5) then true
+ else if eq_vec b__0 ('b"11100" : mword 5) then true
+ else if eq_vec b__0 ('b"11101" : mword 5) then true
+ else if eq_vec b__0 ('b"11110" : mword 5) then true
+ else if eq_vec b__0 ('b"11111" : mword 5) then true
+ else false.
-Definition _update_Ustatus_bits (v : Ustatus) (x : mword 64)
-: Ustatus :=
-
- {[ v with
- Ustatus_Ustatus_chunk_0 :=
- (update_subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
+Definition freg_name_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "ft0" then true
+ else if generic_eq p0_ "ft1" then true
+ else if generic_eq p0_ "ft2" then true
+ else if generic_eq p0_ "ft3" then true
+ else if generic_eq p0_ "ft4" then true
+ else if generic_eq p0_ "ft5" then true
+ else if generic_eq p0_ "ft6" then true
+ else if generic_eq p0_ "ft7" then true
+ else if generic_eq p0_ "fs0" then true
+ else if generic_eq p0_ "fs1" then true
+ else if generic_eq p0_ "fa0" then true
+ else if generic_eq p0_ "fa1" then true
+ else if generic_eq p0_ "fa2" then true
+ else if generic_eq p0_ "fa3" then true
+ else if generic_eq p0_ "fa4" then true
+ else if generic_eq p0_ "fa5" then true
+ else if generic_eq p0_ "fa6" then true
+ else if generic_eq p0_ "fa7" then true
+ else if generic_eq p0_ "fs2" then true
+ else if generic_eq p0_ "fs3" then true
+ else if generic_eq p0_ "fs4" then true
+ else if generic_eq p0_ "fs5" then true
+ else if generic_eq p0_ "fs6" then true
+ else if generic_eq p0_ "fs7" then true
+ else if generic_eq p0_ "fs8" then true
+ else if generic_eq p0_ "fs9" then true
+ else if generic_eq p0_ "fs10" then true
+ else if generic_eq p0_ "fs11" then true
+ else if generic_eq p0_ "ft8" then true
+ else if generic_eq p0_ "ft9" then true
+ else if generic_eq p0_ "ft10" then true
+ else if generic_eq p0_ "ft11" then true
+ else false.
-Definition _get_Ustatus_UPIE (v : Ustatus)
-: mword 1 :=
-
- subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 4 4.
+Definition _s452_ (_s453_ : string) : option string :=
+ let _s454_ := _s453_ in
+ if string_startswith _s454_ "ft11" then
+ match (string_drop _s454_ (projT1 (string_length "ft11"))) with | s_ => Some s_ end
+ else None.
-Definition _set_Ustatus_UPIE (r_ref : register_ref regstate register_value Ustatus) (v : mword 1)
-: M (unit) :=
-
- (reg_deref r_ref) >>= fun r =>
- let r :=
- {[ r with
- Ustatus_Ustatus_chunk_0 :=
- (update_subrange_vec_dec r.(Ustatus_Ustatus_chunk_0) 4 4 (subrange_vec_dec v 0 0)) ]}
- : Ustatus in
- write_reg r_ref r
- : M (unit).
+Definition _s448_ (_s449_ : string) : option string :=
+ let _s450_ := _s449_ in
+ if string_startswith _s450_ "ft10" then
+ match (string_drop _s450_ (projT1 (string_length "ft10"))) with | s_ => Some s_ end
+ else None.
-Definition _update_Ustatus_UPIE (v : Ustatus) (x : mword 1)
-: Ustatus :=
-
- {[ v with
- Ustatus_Ustatus_chunk_0 :=
- (update_subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
+Definition _s444_ (_s445_ : string) : option string :=
+ let _s446_ := _s445_ in
+ if string_startswith _s446_ "ft9" then
+ match (string_drop _s446_ (projT1 (string_length "ft9"))) with | s_ => Some s_ end
+ else None.
-Definition _get_Ustatus_UIE (v : Ustatus)
-: mword 1 :=
-
- subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 0 0.
+Definition _s440_ (_s441_ : string) : option string :=
+ let _s442_ := _s441_ in
+ if string_startswith _s442_ "ft8" then
+ match (string_drop _s442_ (projT1 (string_length "ft8"))) with | s_ => Some s_ end
+ else None.
-Definition _set_Ustatus_UIE (r_ref : register_ref regstate register_value Ustatus) (v : mword 1)
-: M (unit) :=
-
- (reg_deref r_ref) >>= fun r =>
- let r :=
- {[ r with
- Ustatus_Ustatus_chunk_0 :=
- (update_subrange_vec_dec r.(Ustatus_Ustatus_chunk_0) 0 0 (subrange_vec_dec v 0 0)) ]}
- : Ustatus in
- write_reg r_ref r
- : M (unit).
+Definition _s436_ (_s437_ : string) : option string :=
+ let _s438_ := _s437_ in
+ if string_startswith _s438_ "fs11" then
+ match (string_drop _s438_ (projT1 (string_length "fs11"))) with | s_ => Some s_ end
+ else None.
-Definition _update_Ustatus_UIE (v : Ustatus) (x : mword 1)
-: Ustatus :=
-
- {[ v with
- Ustatus_Ustatus_chunk_0 :=
- (update_subrange_vec_dec v.(Ustatus_Ustatus_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
+Definition _s432_ (_s433_ : string) : option string :=
+ let _s434_ := _s433_ in
+ if string_startswith _s434_ "fs10" then
+ match (string_drop _s434_ (projT1 (string_length "fs10"))) with | s_ => Some s_ end
+ else None.
-Definition lower_sstatus (s : Sstatus)
-: Ustatus :=
-
- let u := Mk_Ustatus (EXTZ 64 (vec_of_bits [B0] : mword 1)) in
- let u := _update_Ustatus_UPIE u (_get_Sstatus_UPIE s) in
- _update_Ustatus_UIE u (_get_Sstatus_UIE s).
+Definition _s428_ (_s429_ : string) : option string :=
+ let _s430_ := _s429_ in
+ if string_startswith _s430_ "fs9" then
+ match (string_drop _s430_ (projT1 (string_length "fs9"))) with | s_ => Some s_ end
+ else None.
-Definition lift_ustatus (s : Sstatus) (u : Ustatus)
-: Sstatus :=
-
- let s := _update_Sstatus_UPIE s (_get_Ustatus_UPIE u) in
- _update_Sstatus_UIE s (_get_Ustatus_UIE u).
+Definition _s424_ (_s425_ : string) : option string :=
+ let _s426_ := _s425_ in
+ if string_startswith _s426_ "fs8" then
+ match (string_drop _s426_ (projT1 (string_length "fs8"))) with | s_ => Some s_ end
+ else None.
-Definition legalize_ustatus (m : Mstatus) (v : mword 64)
-: M (Mstatus) :=
-
- let u := Mk_Ustatus v in
- let s := lower_mstatus m in
- let s := lift_ustatus s u in
- (lift_sstatus m s)
- : M (Mstatus).
+Definition _s420_ (_s421_ : string) : option string :=
+ let _s422_ := _s421_ in
+ if string_startswith _s422_ "fs7" then
+ match (string_drop _s422_ (projT1 (string_length "fs7"))) with | s_ => Some s_ end
+ else None.
-Definition Mk_Uinterrupts (v : mword 64)
-: Uinterrupts :=
-
- {| Uinterrupts_Uinterrupts_chunk_0 := (subrange_vec_dec v 63 0) |}.
+Definition _s416_ (_s417_ : string) : option string :=
+ let _s418_ := _s417_ in
+ if string_startswith _s418_ "fs6" then
+ match (string_drop _s418_ (projT1 (string_length "fs6"))) with | s_ => Some s_ end
+ else None.
-Definition _get_Uinterrupts_bits (v : Uinterrupts)
-: mword 64 :=
-
- subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 63 0.
+Definition _s412_ (_s413_ : string) : option string :=
+ let _s414_ := _s413_ in
+ if string_startswith _s414_ "fs5" then
+ match (string_drop _s414_ (projT1 (string_length "fs5"))) with | s_ => Some s_ end
+ else None.
-Definition _set_Uinterrupts_bits
-(r_ref : register_ref regstate register_value Uinterrupts) (v : mword 64)
-: M (unit) :=
-
- (reg_deref r_ref) >>= fun r =>
- let r :=
- {[ r with
- Uinterrupts_Uinterrupts_chunk_0 :=
- (update_subrange_vec_dec r.(Uinterrupts_Uinterrupts_chunk_0) 63 0 (subrange_vec_dec v 63 0)) ]}
- : Uinterrupts in
- write_reg r_ref r
- : M (unit).
+Definition _s408_ (_s409_ : string) : option string :=
+ let _s410_ := _s409_ in
+ if string_startswith _s410_ "fs4" then
+ match (string_drop _s410_ (projT1 (string_length "fs4"))) with | s_ => Some s_ end
+ else None.
-Definition _update_Uinterrupts_bits (v : Uinterrupts) (x : mword 64)
-: Uinterrupts :=
-
- {[ v with
- Uinterrupts_Uinterrupts_chunk_0 :=
- (update_subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
+Definition _s404_ (_s405_ : string) : option string :=
+ let _s406_ := _s405_ in
+ if string_startswith _s406_ "fs3" then
+ match (string_drop _s406_ (projT1 (string_length "fs3"))) with | s_ => Some s_ end
+ else None.
-Definition _get_Uinterrupts_UEI (v : Uinterrupts)
-: mword 1 :=
-
- subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 8 8.
+Definition _s400_ (_s401_ : string) : option string :=
+ let _s402_ := _s401_ in
+ if string_startswith _s402_ "fs2" then
+ match (string_drop _s402_ (projT1 (string_length "fs2"))) with | s_ => Some s_ end
+ else None.
-Definition _set_Uinterrupts_UEI
-(r_ref : register_ref regstate register_value Uinterrupts) (v : mword 1)
+Definition _s396_ (_s397_ : string) : option string :=
+ let _s398_ := _s397_ in
+ if string_startswith _s398_ "fa7" then
+ match (string_drop _s398_ (projT1 (string_length "fa7"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s392_ (_s393_ : string) : option string :=
+ let _s394_ := _s393_ in
+ if string_startswith _s394_ "fa6" then
+ match (string_drop _s394_ (projT1 (string_length "fa6"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s388_ (_s389_ : string) : option string :=
+ let _s390_ := _s389_ in
+ if string_startswith _s390_ "fa5" then
+ match (string_drop _s390_ (projT1 (string_length "fa5"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s384_ (_s385_ : string) : option string :=
+ let _s386_ := _s385_ in
+ if string_startswith _s386_ "fa4" then
+ match (string_drop _s386_ (projT1 (string_length "fa4"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s380_ (_s381_ : string) : option string :=
+ let _s382_ := _s381_ in
+ if string_startswith _s382_ "fa3" then
+ match (string_drop _s382_ (projT1 (string_length "fa3"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s376_ (_s377_ : string) : option string :=
+ let _s378_ := _s377_ in
+ if string_startswith _s378_ "fa2" then
+ match (string_drop _s378_ (projT1 (string_length "fa2"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s372_ (_s373_ : string) : option string :=
+ let _s374_ := _s373_ in
+ if string_startswith _s374_ "fa1" then
+ match (string_drop _s374_ (projT1 (string_length "fa1"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s368_ (_s369_ : string) : option string :=
+ let _s370_ := _s369_ in
+ if string_startswith _s370_ "fa0" then
+ match (string_drop _s370_ (projT1 (string_length "fa0"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s364_ (_s365_ : string) : option string :=
+ let _s366_ := _s365_ in
+ if string_startswith _s366_ "fs1" then
+ match (string_drop _s366_ (projT1 (string_length "fs1"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s360_ (_s361_ : string) : option string :=
+ let _s362_ := _s361_ in
+ if string_startswith _s362_ "fs0" then
+ match (string_drop _s362_ (projT1 (string_length "fs0"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s356_ (_s357_ : string) : option string :=
+ let _s358_ := _s357_ in
+ if string_startswith _s358_ "ft7" then
+ match (string_drop _s358_ (projT1 (string_length "ft7"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s352_ (_s353_ : string) : option string :=
+ let _s354_ := _s353_ in
+ if string_startswith _s354_ "ft6" then
+ match (string_drop _s354_ (projT1 (string_length "ft6"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s348_ (_s349_ : string) : option string :=
+ let _s350_ := _s349_ in
+ if string_startswith _s350_ "ft5" then
+ match (string_drop _s350_ (projT1 (string_length "ft5"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s344_ (_s345_ : string) : option string :=
+ let _s346_ := _s345_ in
+ if string_startswith _s346_ "ft4" then
+ match (string_drop _s346_ (projT1 (string_length "ft4"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s340_ (_s341_ : string) : option string :=
+ let _s342_ := _s341_ in
+ if string_startswith _s342_ "ft3" then
+ match (string_drop _s342_ (projT1 (string_length "ft3"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s336_ (_s337_ : string) : option string :=
+ let _s338_ := _s337_ in
+ if string_startswith _s338_ "ft2" then
+ match (string_drop _s338_ (projT1 (string_length "ft2"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s332_ (_s333_ : string) : option string :=
+ let _s334_ := _s333_ in
+ if string_startswith _s334_ "ft1" then
+ match (string_drop _s334_ (projT1 (string_length "ft1"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s328_ (_s329_ : string) : option string :=
+ let _s330_ := _s329_ in
+ if string_startswith _s330_ "ft0" then
+ match (string_drop _s330_ (projT1 (string_length "ft0"))) with | s_ => Some s_ end
+ else None.
+
+Definition freg_name_matches_prefix (arg_ : string)
+: M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s331_ := arg_ in
+ (if match (_s328_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s328_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"00000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s332_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s332_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"00001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s336_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s336_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"00010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s340_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s340_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"00011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s344_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s344_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"00100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s348_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s348_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"00101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s352_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s352_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"00110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s356_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s356_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"00111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s360_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s360_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"01000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s364_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s364_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"01001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s368_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s368_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"01010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s372_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s372_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"01011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s376_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s376_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"01100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s380_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s380_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"01101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s384_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s384_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"01110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s388_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s388_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"01111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s392_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s392_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"10000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s396_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s396_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"10001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s400_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s400_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"10010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s404_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s404_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"10011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s408_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s408_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"10100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s412_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s412_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"10101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s416_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s416_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"10110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s420_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s420_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"10111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s424_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s424_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"11000"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s428_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s428_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"11001"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s432_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s432_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"11010"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s436_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s436_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"11011"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s440_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s440_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"11100"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s444_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s444_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"11101"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s448_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s448_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"11110"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s452_ _s331_) with | Some s_ => true | _ => false end then
+ (match (_s452_ _s331_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"11111"
+ : mword 5, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 5 * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition init_fdext_regs '(tt : unit) : M (unit) :=
+ write_reg f0_ref zero_freg >>
+ write_reg f1_ref zero_freg >>
+ write_reg f2_ref zero_freg >>
+ write_reg f3_ref zero_freg >>
+ write_reg f4_ref zero_freg >>
+ write_reg f5_ref zero_freg >>
+ write_reg f6_ref zero_freg >>
+ write_reg f7_ref zero_freg >>
+ write_reg f8_ref zero_freg >>
+ write_reg f9_ref zero_freg >>
+ write_reg f10_ref zero_freg >>
+ write_reg f11_ref zero_freg >>
+ write_reg f12_ref zero_freg >>
+ write_reg f13_ref zero_freg >>
+ write_reg f14_ref zero_freg >>
+ write_reg f15_ref zero_freg >>
+ write_reg f16_ref zero_freg >>
+ write_reg f17_ref zero_freg >>
+ write_reg f18_ref zero_freg >>
+ write_reg f19_ref zero_freg >>
+ write_reg f20_ref zero_freg >>
+ write_reg f21_ref zero_freg >>
+ write_reg f22_ref zero_freg >>
+ write_reg f23_ref zero_freg >>
+ write_reg f24_ref zero_freg >>
+ write_reg f25_ref zero_freg >>
+ write_reg f26_ref zero_freg >>
+ write_reg f27_ref zero_freg >>
+ write_reg f28_ref zero_freg >>
+ write_reg f29_ref zero_freg >>
+ write_reg f30_ref zero_freg >> write_reg f31_ref zero_freg : M (unit).
+
+Definition Mk_Fcsr (v : mword 32) : Fcsr := {| Fcsr_Fcsr_chunk_0 := (subrange_vec_dec v 31 0) |}.
+
+Definition _get_Fcsr_bits (v : Fcsr) : mword 32 := subrange_vec_dec v.(Fcsr_Fcsr_chunk_0) 31 0.
+
+Definition _set_Fcsr_bits (r_ref : register_ref regstate register_value Fcsr) (v : mword 32)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
- Uinterrupts_Uinterrupts_chunk_0 :=
- (update_subrange_vec_dec r.(Uinterrupts_Uinterrupts_chunk_0) 8 8 (subrange_vec_dec v 0 0)) ]}
- : Uinterrupts in
+ Fcsr_Fcsr_chunk_0 :=
+ (update_subrange_vec_dec r.(Fcsr_Fcsr_chunk_0) 31 0 (subrange_vec_dec v 31 0)) ]}
+ : Fcsr in
write_reg r_ref r
: M (unit).
-Definition _update_Uinterrupts_UEI (v : Uinterrupts) (x : mword 1)
-: Uinterrupts :=
-
+Definition _update_Fcsr_bits (v : Fcsr) (x : mword 32) : Fcsr :=
{[ v with
- Uinterrupts_Uinterrupts_chunk_0 :=
- (update_subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 8 8 (subrange_vec_dec x 0 0)) ]}.
+ Fcsr_Fcsr_chunk_0 :=
+ (update_subrange_vec_dec v.(Fcsr_Fcsr_chunk_0) 31 0 (subrange_vec_dec x 31 0)) ]}.
-Definition _get_Uinterrupts_UTI (v : Uinterrupts)
-: mword 1 :=
-
- subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 4 4.
+Definition _get_Fcsr_FRM (v : Fcsr) : mword 3 := subrange_vec_dec v.(Fcsr_Fcsr_chunk_0) 7 5.
-Definition _set_Uinterrupts_UTI
-(r_ref : register_ref regstate register_value Uinterrupts) (v : mword 1)
+Definition _set_Fcsr_FRM (r_ref : register_ref regstate register_value Fcsr) (v : mword 3)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
- Uinterrupts_Uinterrupts_chunk_0 :=
- (update_subrange_vec_dec r.(Uinterrupts_Uinterrupts_chunk_0) 4 4 (subrange_vec_dec v 0 0)) ]}
- : Uinterrupts in
+ Fcsr_Fcsr_chunk_0 :=
+ (update_subrange_vec_dec r.(Fcsr_Fcsr_chunk_0) 7 5 (subrange_vec_dec v 2 0)) ]}
+ : Fcsr in
write_reg r_ref r
: M (unit).
-Definition _update_Uinterrupts_UTI (v : Uinterrupts) (x : mword 1)
-: Uinterrupts :=
-
+Definition _update_Fcsr_FRM (v : Fcsr) (x : mword 3) : Fcsr :=
{[ v with
- Uinterrupts_Uinterrupts_chunk_0 :=
- (update_subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
+ Fcsr_Fcsr_chunk_0 :=
+ (update_subrange_vec_dec v.(Fcsr_Fcsr_chunk_0) 7 5 (subrange_vec_dec x 2 0)) ]}.
-Definition _get_Uinterrupts_USI (v : Uinterrupts)
-: mword 1 :=
-
- subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 0 0.
+Definition _get_Fcsr_FFLAGS (v : Fcsr) : mword 5 := subrange_vec_dec v.(Fcsr_Fcsr_chunk_0) 4 0.
-Definition _set_Uinterrupts_USI
-(r_ref : register_ref regstate register_value Uinterrupts) (v : mword 1)
+Definition _set_Fcsr_FFLAGS (r_ref : register_ref regstate register_value Fcsr) (v : mword 5)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
- Uinterrupts_Uinterrupts_chunk_0 :=
- (update_subrange_vec_dec r.(Uinterrupts_Uinterrupts_chunk_0) 0 0 (subrange_vec_dec v 0 0)) ]}
- : Uinterrupts in
+ Fcsr_Fcsr_chunk_0 :=
+ (update_subrange_vec_dec r.(Fcsr_Fcsr_chunk_0) 4 0 (subrange_vec_dec v 4 0)) ]}
+ : Fcsr in
write_reg r_ref r
: M (unit).
-Definition _update_Uinterrupts_USI (v : Uinterrupts) (x : mword 1)
-: Uinterrupts :=
-
+Definition _update_Fcsr_FFLAGS (v : Fcsr) (x : mword 5) : Fcsr :=
{[ v with
- Uinterrupts_Uinterrupts_chunk_0 :=
- (update_subrange_vec_dec v.(Uinterrupts_Uinterrupts_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
+ Fcsr_Fcsr_chunk_0 :=
+ (update_subrange_vec_dec v.(Fcsr_Fcsr_chunk_0) 4 0 (subrange_vec_dec x 4 0)) ]}.
+
+Definition ext_write_fcsr (frm : mword 3) (fflags : mword 5) : M (unit) :=
+ (_set_Fcsr_FRM fcsr_ref frm) >>
+ (_set_Fcsr_FFLAGS fcsr_ref fflags) >>
+ (update_softfloat_fflags fflags) >> (dirty_fd_context tt) : M (unit).
+
+Definition write_fflags (fflags : mword 5) : M (unit) :=
+ read_reg fcsr_ref >>= fun w__0 : Fcsr =>
+ (if neq_vec (_get_Fcsr_FFLAGS w__0) fflags then (dirty_fd_context tt) : M (unit)
+ else returnm tt) >>
+ (_set_Fcsr_FFLAGS fcsr_ref fflags)
+ : M (unit).
-Definition lower_sip (s : Sinterrupts) (d : Sinterrupts)
-: Uinterrupts :=
-
- let u : Uinterrupts := Mk_Uinterrupts (EXTZ 64 (vec_of_bits [B0] : mword 1)) in
- let u := _update_Uinterrupts_UEI u (and_vec (_get_Sinterrupts_UEI s) (_get_Sinterrupts_UEI d)) in
- let u := _update_Uinterrupts_UTI u (and_vec (_get_Sinterrupts_UTI s) (_get_Sinterrupts_UTI d)) in
- _update_Uinterrupts_USI u (and_vec (_get_Sinterrupts_USI s) (_get_Sinterrupts_USI d)).
+Definition accrue_fflags (flags : mword 5) : M (unit) :=
+ read_reg fcsr_ref >>= fun w__0 : Fcsr =>
+ let f := or_vec (_get_Fcsr_FFLAGS w__0) flags in
+ read_reg fcsr_ref >>= fun w__1 : Fcsr =>
+ (if neq_vec (_get_Fcsr_FFLAGS w__1) f then
+ (_set_Fcsr_FFLAGS fcsr_ref f) >>
+ (update_softfloat_fflags f) >> (dirty_fd_context tt) : M (unit)
+ else returnm tt)
+ : M (unit).
-Definition lower_sie (s : Sinterrupts) (d : Sinterrupts)
-: Uinterrupts :=
-
- let u : Uinterrupts := Mk_Uinterrupts (EXTZ 64 (vec_of_bits [B0] : mword 1)) in
- let u := _update_Uinterrupts_UEI u (and_vec (_get_Sinterrupts_UEI s) (_get_Sinterrupts_UEI d)) in
- let u := _update_Uinterrupts_UTI u (and_vec (_get_Sinterrupts_UTI s) (_get_Sinterrupts_UTI d)) in
- _update_Uinterrupts_USI u (and_vec (_get_Sinterrupts_USI s) (_get_Sinterrupts_USI d)).
+Definition csr_name_map_forwards (arg_ : mword 12) : string :=
+ let b__0 := arg_ in
+ if eq_vec b__0 (Ox"000" : mword 12) then "ustatus"
+ else if eq_vec b__0 (Ox"004" : mword 12) then "uie"
+ else if eq_vec b__0 (Ox"005" : mword 12) then "utvec"
+ else if eq_vec b__0 (Ox"040" : mword 12) then "uscratch"
+ else if eq_vec b__0 (Ox"041" : mword 12) then "uepc"
+ else if eq_vec b__0 (Ox"042" : mword 12) then "ucause"
+ else if eq_vec b__0 (Ox"043" : mword 12) then "utval"
+ else if eq_vec b__0 (Ox"044" : mword 12) then "uip"
+ else if eq_vec b__0 (Ox"001" : mword 12) then "fflags"
+ else if eq_vec b__0 (Ox"002" : mword 12) then "frm"
+ else if eq_vec b__0 (Ox"003" : mword 12) then "fcsr"
+ else if eq_vec b__0 (Ox"C00" : mword 12) then "cycle"
+ else if eq_vec b__0 (Ox"C01" : mword 12) then "time"
+ else if eq_vec b__0 (Ox"C02" : mword 12) then "instret"
+ else if eq_vec b__0 (Ox"C80" : mword 12) then "cycleh"
+ else if eq_vec b__0 (Ox"C81" : mword 12) then "timeh"
+ else if eq_vec b__0 (Ox"C82" : mword 12) then "instreth"
+ else if eq_vec b__0 (Ox"100" : mword 12) then "sstatus"
+ else if eq_vec b__0 (Ox"102" : mword 12) then "sedeleg"
+ else if eq_vec b__0 (Ox"103" : mword 12) then "sideleg"
+ else if eq_vec b__0 (Ox"104" : mword 12) then "sie"
+ else if eq_vec b__0 (Ox"105" : mword 12) then "stvec"
+ else if eq_vec b__0 (Ox"106" : mword 12) then "scounteren"
+ else if eq_vec b__0 (Ox"140" : mword 12) then "sscratch"
+ else if eq_vec b__0 (Ox"141" : mword 12) then "sepc"
+ else if eq_vec b__0 (Ox"142" : mword 12) then "scause"
+ else if eq_vec b__0 (Ox"143" : mword 12) then "stval"
+ else if eq_vec b__0 (Ox"144" : mword 12) then "sip"
+ else if eq_vec b__0 (Ox"180" : mword 12) then "satp"
+ else if eq_vec b__0 (Ox"F11" : mword 12) then "mvendorid"
+ else if eq_vec b__0 (Ox"F12" : mword 12) then "marchid"
+ else if eq_vec b__0 (Ox"F13" : mword 12) then "mimpid"
+ else if eq_vec b__0 (Ox"F14" : mword 12) then "mhartid"
+ else if eq_vec b__0 (Ox"300" : mword 12) then "mstatus"
+ else if eq_vec b__0 (Ox"301" : mword 12) then "misa"
+ else if eq_vec b__0 (Ox"302" : mword 12) then "medeleg"
+ else if eq_vec b__0 (Ox"303" : mword 12) then "mideleg"
+ else if eq_vec b__0 (Ox"304" : mword 12) then "mie"
+ else if eq_vec b__0 (Ox"305" : mword 12) then "mtvec"
+ else if eq_vec b__0 (Ox"306" : mword 12) then "mcounteren"
+ else if eq_vec b__0 (Ox"340" : mword 12) then "mscratch"
+ else if eq_vec b__0 (Ox"341" : mword 12) then "mepc"
+ else if eq_vec b__0 (Ox"342" : mword 12) then "mcause"
+ else if eq_vec b__0 (Ox"343" : mword 12) then "mtval"
+ else if eq_vec b__0 (Ox"344" : mword 12) then "mip"
+ else if eq_vec b__0 (Ox"3A0" : mword 12) then "pmpcfg0"
+ else if eq_vec b__0 (Ox"3A1" : mword 12) then "pmpcfg1"
+ else if eq_vec b__0 (Ox"3A2" : mword 12) then "pmpcfg2"
+ else if eq_vec b__0 (Ox"3A3" : mword 12) then "pmpcfg3"
+ else if eq_vec b__0 (Ox"3B0" : mword 12) then "pmpaddr0"
+ else if eq_vec b__0 (Ox"3B1" : mword 12) then "pmpaddr1"
+ else if eq_vec b__0 (Ox"3B2" : mword 12) then "pmpaddr2"
+ else if eq_vec b__0 (Ox"3B3" : mword 12) then "pmpaddr3"
+ else if eq_vec b__0 (Ox"3B4" : mword 12) then "pmpaddr4"
+ else if eq_vec b__0 (Ox"3B5" : mword 12) then "pmpaddr5"
+ else if eq_vec b__0 (Ox"3B6" : mword 12) then "pmpaddr6"
+ else if eq_vec b__0 (Ox"3B7" : mword 12) then "pmpaddr7"
+ else if eq_vec b__0 (Ox"3B8" : mword 12) then "pmpaddr8"
+ else if eq_vec b__0 (Ox"3B9" : mword 12) then "pmpaddr9"
+ else if eq_vec b__0 (Ox"3BA" : mword 12) then "pmpaddr10"
+ else if eq_vec b__0 (Ox"3BB" : mword 12) then "pmpaddr11"
+ else if eq_vec b__0 (Ox"3BC" : mword 12) then "pmpaddr12"
+ else if eq_vec b__0 (Ox"3BD" : mword 12) then "pmpaddr13"
+ else if eq_vec b__0 (Ox"3BE" : mword 12) then "pmpaddr14"
+ else if eq_vec b__0 (Ox"3BF" : mword 12) then "pmpaddr15"
+ else if eq_vec b__0 (Ox"B00" : mword 12) then "mcycle"
+ else if eq_vec b__0 (Ox"B02" : mword 12) then "minstret"
+ else if eq_vec b__0 (Ox"B80" : mword 12) then "mcycleh"
+ else if eq_vec b__0 (Ox"B82" : mword 12) then "minstreth"
+ else if eq_vec b__0 (Ox"7A0" : mword 12) then "tselect"
+ else if eq_vec b__0 (Ox"7A1" : mword 12) then "tdata1"
+ else if eq_vec b__0 (Ox"7A2" : mword 12) then "tdata2"
+ else if eq_vec b__0 (Ox"7A3" : mword 12) then "tdata3"
+ else decimal_string_of_bits b__0.
+
+Definition csr_name_map_backwards (arg_ : string) : M (mword 12) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "ustatus" then returnm (Ox"000" : mword 12)
+ else if generic_eq p0_ "uie" then returnm (Ox"004" : mword 12)
+ else if generic_eq p0_ "utvec" then returnm (Ox"005" : mword 12)
+ else if generic_eq p0_ "uscratch" then returnm (Ox"040" : mword 12)
+ else if generic_eq p0_ "uepc" then returnm (Ox"041" : mword 12)
+ else if generic_eq p0_ "ucause" then returnm (Ox"042" : mword 12)
+ else if generic_eq p0_ "utval" then returnm (Ox"043" : mword 12)
+ else if generic_eq p0_ "uip" then returnm (Ox"044" : mword 12)
+ else if generic_eq p0_ "fflags" then returnm (Ox"001" : mword 12)
+ else if generic_eq p0_ "frm" then returnm (Ox"002" : mword 12)
+ else if generic_eq p0_ "fcsr" then returnm (Ox"003" : mword 12)
+ else if generic_eq p0_ "cycle" then returnm (Ox"C00" : mword 12)
+ else if generic_eq p0_ "time" then returnm (Ox"C01" : mword 12)
+ else if generic_eq p0_ "instret" then returnm (Ox"C02" : mword 12)
+ else if generic_eq p0_ "cycleh" then returnm (Ox"C80" : mword 12)
+ else if generic_eq p0_ "timeh" then returnm (Ox"C81" : mword 12)
+ else if generic_eq p0_ "instreth" then returnm (Ox"C82" : mword 12)
+ else if generic_eq p0_ "sstatus" then returnm (Ox"100" : mword 12)
+ else if generic_eq p0_ "sedeleg" then returnm (Ox"102" : mword 12)
+ else if generic_eq p0_ "sideleg" then returnm (Ox"103" : mword 12)
+ else if generic_eq p0_ "sie" then returnm (Ox"104" : mword 12)
+ else if generic_eq p0_ "stvec" then returnm (Ox"105" : mword 12)
+ else if generic_eq p0_ "scounteren" then returnm (Ox"106" : mword 12)
+ else if generic_eq p0_ "sscratch" then returnm (Ox"140" : mword 12)
+ else if generic_eq p0_ "sepc" then returnm (Ox"141" : mword 12)
+ else if generic_eq p0_ "scause" then returnm (Ox"142" : mword 12)
+ else if generic_eq p0_ "stval" then returnm (Ox"143" : mword 12)
+ else if generic_eq p0_ "sip" then returnm (Ox"144" : mword 12)
+ else if generic_eq p0_ "satp" then returnm (Ox"180" : mword 12)
+ else if generic_eq p0_ "mvendorid" then returnm (Ox"F11" : mword 12)
+ else if generic_eq p0_ "marchid" then returnm (Ox"F12" : mword 12)
+ else if generic_eq p0_ "mimpid" then returnm (Ox"F13" : mword 12)
+ else if generic_eq p0_ "mhartid" then returnm (Ox"F14" : mword 12)
+ else if generic_eq p0_ "mstatus" then returnm (Ox"300" : mword 12)
+ else if generic_eq p0_ "misa" then returnm (Ox"301" : mword 12)
+ else if generic_eq p0_ "medeleg" then returnm (Ox"302" : mword 12)
+ else if generic_eq p0_ "mideleg" then returnm (Ox"303" : mword 12)
+ else if generic_eq p0_ "mie" then returnm (Ox"304" : mword 12)
+ else if generic_eq p0_ "mtvec" then returnm (Ox"305" : mword 12)
+ else if generic_eq p0_ "mcounteren" then returnm (Ox"306" : mword 12)
+ else if generic_eq p0_ "mscratch" then returnm (Ox"340" : mword 12)
+ else if generic_eq p0_ "mepc" then returnm (Ox"341" : mword 12)
+ else if generic_eq p0_ "mcause" then returnm (Ox"342" : mword 12)
+ else if generic_eq p0_ "mtval" then returnm (Ox"343" : mword 12)
+ else if generic_eq p0_ "mip" then returnm (Ox"344" : mword 12)
+ else if generic_eq p0_ "pmpcfg0" then returnm (Ox"3A0" : mword 12)
+ else if generic_eq p0_ "pmpcfg1" then returnm (Ox"3A1" : mword 12)
+ else if generic_eq p0_ "pmpcfg2" then returnm (Ox"3A2" : mword 12)
+ else if generic_eq p0_ "pmpcfg3" then returnm (Ox"3A3" : mword 12)
+ else if generic_eq p0_ "pmpaddr0" then returnm (Ox"3B0" : mword 12)
+ else if generic_eq p0_ "pmpaddr1" then returnm (Ox"3B1" : mword 12)
+ else if generic_eq p0_ "pmpaddr2" then returnm (Ox"3B2" : mword 12)
+ else if generic_eq p0_ "pmpaddr3" then returnm (Ox"3B3" : mword 12)
+ else if generic_eq p0_ "pmpaddr4" then returnm (Ox"3B4" : mword 12)
+ else if generic_eq p0_ "pmpaddr5" then returnm (Ox"3B5" : mword 12)
+ else if generic_eq p0_ "pmpaddr6" then returnm (Ox"3B6" : mword 12)
+ else if generic_eq p0_ "pmpaddr7" then returnm (Ox"3B7" : mword 12)
+ else if generic_eq p0_ "pmpaddr8" then returnm (Ox"3B8" : mword 12)
+ else if generic_eq p0_ "pmpaddr9" then returnm (Ox"3B9" : mword 12)
+ else if generic_eq p0_ "pmpaddr10" then returnm (Ox"3BA" : mword 12)
+ else if generic_eq p0_ "pmpaddr11" then returnm (Ox"3BB" : mword 12)
+ else if generic_eq p0_ "pmpaddr12" then returnm (Ox"3BC" : mword 12)
+ else if generic_eq p0_ "pmpaddr13" then returnm (Ox"3BD" : mword 12)
+ else if generic_eq p0_ "pmpaddr14" then returnm (Ox"3BE" : mword 12)
+ else if generic_eq p0_ "pmpaddr15" then returnm (Ox"3BF" : mword 12)
+ else if generic_eq p0_ "mcycle" then returnm (Ox"B00" : mword 12)
+ else if generic_eq p0_ "minstret" then returnm (Ox"B02" : mword 12)
+ else if generic_eq p0_ "mcycleh" then returnm (Ox"B80" : mword 12)
+ else if generic_eq p0_ "minstreth" then returnm (Ox"B82" : mword 12)
+ else if generic_eq p0_ "tselect" then returnm (Ox"7A0" : mword 12)
+ else if generic_eq p0_ "tdata1" then returnm (Ox"7A1" : mword 12)
+ else if generic_eq p0_ "tdata2" then returnm (Ox"7A2" : mword 12)
+ else if generic_eq p0_ "tdata3" then returnm (Ox"7A3" : mword 12)
+ else
+ (and_boolM (returnm ((hex_bits_12_backwards_matches p0_) : bool))
+ ((if hex_bits_12_backwards_matches p0_ then
+ (hex_bits_12_backwards p0_) >>= fun reg => returnm (true : bool)
+ else returnm false)
+ : M (bool))) >>= fun w__1 : bool =>
+ (if sumbool_of_bool w__1 then hex_bits_12_backwards p0_
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 12))
+ : M (mword 12).
-Definition lift_uip (o : Sinterrupts) (d : Sinterrupts) (u : Uinterrupts)
-: Sinterrupts :=
-
- let s : Sinterrupts := o in
- if ((eq_vec (_get_Sinterrupts_USI d) ((bool_to_bits true) : mword 1))) then
- _update_Sinterrupts_USI s (_get_Uinterrupts_USI u)
- else s.
+Definition csr_name_map_forwards_matches (arg_ : mword 12) : bool :=
+ let b__0 := arg_ in
+ if eq_vec b__0 (Ox"000" : mword 12) then true
+ else if eq_vec b__0 (Ox"004" : mword 12) then true
+ else if eq_vec b__0 (Ox"005" : mword 12) then true
+ else if eq_vec b__0 (Ox"040" : mword 12) then true
+ else if eq_vec b__0 (Ox"041" : mword 12) then true
+ else if eq_vec b__0 (Ox"042" : mword 12) then true
+ else if eq_vec b__0 (Ox"043" : mword 12) then true
+ else if eq_vec b__0 (Ox"044" : mword 12) then true
+ else if eq_vec b__0 (Ox"001" : mword 12) then true
+ else if eq_vec b__0 (Ox"002" : mword 12) then true
+ else if eq_vec b__0 (Ox"003" : mword 12) then true
+ else if eq_vec b__0 (Ox"C00" : mword 12) then true
+ else if eq_vec b__0 (Ox"C01" : mword 12) then true
+ else if eq_vec b__0 (Ox"C02" : mword 12) then true
+ else if eq_vec b__0 (Ox"C80" : mword 12) then true
+ else if eq_vec b__0 (Ox"C81" : mword 12) then true
+ else if eq_vec b__0 (Ox"C82" : mword 12) then true
+ else if eq_vec b__0 (Ox"100" : mword 12) then true
+ else if eq_vec b__0 (Ox"102" : mword 12) then true
+ else if eq_vec b__0 (Ox"103" : mword 12) then true
+ else if eq_vec b__0 (Ox"104" : mword 12) then true
+ else if eq_vec b__0 (Ox"105" : mword 12) then true
+ else if eq_vec b__0 (Ox"106" : mword 12) then true
+ else if eq_vec b__0 (Ox"140" : mword 12) then true
+ else if eq_vec b__0 (Ox"141" : mword 12) then true
+ else if eq_vec b__0 (Ox"142" : mword 12) then true
+ else if eq_vec b__0 (Ox"143" : mword 12) then true
+ else if eq_vec b__0 (Ox"144" : mword 12) then true
+ else if eq_vec b__0 (Ox"180" : mword 12) then true
+ else if eq_vec b__0 (Ox"F11" : mword 12) then true
+ else if eq_vec b__0 (Ox"F12" : mword 12) then true
+ else if eq_vec b__0 (Ox"F13" : mword 12) then true
+ else if eq_vec b__0 (Ox"F14" : mword 12) then true
+ else if eq_vec b__0 (Ox"300" : mword 12) then true
+ else if eq_vec b__0 (Ox"301" : mword 12) then true
+ else if eq_vec b__0 (Ox"302" : mword 12) then true
+ else if eq_vec b__0 (Ox"303" : mword 12) then true
+ else if eq_vec b__0 (Ox"304" : mword 12) then true
+ else if eq_vec b__0 (Ox"305" : mword 12) then true
+ else if eq_vec b__0 (Ox"306" : mword 12) then true
+ else if eq_vec b__0 (Ox"340" : mword 12) then true
+ else if eq_vec b__0 (Ox"341" : mword 12) then true
+ else if eq_vec b__0 (Ox"342" : mword 12) then true
+ else if eq_vec b__0 (Ox"343" : mword 12) then true
+ else if eq_vec b__0 (Ox"344" : mword 12) then true
+ else if eq_vec b__0 (Ox"3A0" : mword 12) then true
+ else if eq_vec b__0 (Ox"3A1" : mword 12) then true
+ else if eq_vec b__0 (Ox"3A2" : mword 12) then true
+ else if eq_vec b__0 (Ox"3A3" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B0" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B1" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B2" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B3" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B4" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B5" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B6" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B7" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B8" : mword 12) then true
+ else if eq_vec b__0 (Ox"3B9" : mword 12) then true
+ else if eq_vec b__0 (Ox"3BA" : mword 12) then true
+ else if eq_vec b__0 (Ox"3BB" : mword 12) then true
+ else if eq_vec b__0 (Ox"3BC" : mword 12) then true
+ else if eq_vec b__0 (Ox"3BD" : mword 12) then true
+ else if eq_vec b__0 (Ox"3BE" : mword 12) then true
+ else if eq_vec b__0 (Ox"3BF" : mword 12) then true
+ else if eq_vec b__0 (Ox"B00" : mword 12) then true
+ else if eq_vec b__0 (Ox"B02" : mword 12) then true
+ else if eq_vec b__0 (Ox"B80" : mword 12) then true
+ else if eq_vec b__0 (Ox"B82" : mword 12) then true
+ else if eq_vec b__0 (Ox"7A0" : mword 12) then true
+ else if eq_vec b__0 (Ox"7A1" : mword 12) then true
+ else if eq_vec b__0 (Ox"7A2" : mword 12) then true
+ else if eq_vec b__0 (Ox"7A3" : mword 12) then true
+ else true.
-Definition legalize_uip (s : Sinterrupts) (d : Sinterrupts) (v : mword 64)
-: Sinterrupts :=
-
- lift_uip s d (Mk_Uinterrupts v).
+Definition csr_name_map_backwards_matches (arg_ : string) : M (bool) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "ustatus" then returnm true
+ else if generic_eq p0_ "uie" then returnm true
+ else if generic_eq p0_ "utvec" then returnm true
+ else if generic_eq p0_ "uscratch" then returnm true
+ else if generic_eq p0_ "uepc" then returnm true
+ else if generic_eq p0_ "ucause" then returnm true
+ else if generic_eq p0_ "utval" then returnm true
+ else if generic_eq p0_ "uip" then returnm true
+ else if generic_eq p0_ "fflags" then returnm true
+ else if generic_eq p0_ "frm" then returnm true
+ else if generic_eq p0_ "fcsr" then returnm true
+ else if generic_eq p0_ "cycle" then returnm true
+ else if generic_eq p0_ "time" then returnm true
+ else if generic_eq p0_ "instret" then returnm true
+ else if generic_eq p0_ "cycleh" then returnm true
+ else if generic_eq p0_ "timeh" then returnm true
+ else if generic_eq p0_ "instreth" then returnm true
+ else if generic_eq p0_ "sstatus" then returnm true
+ else if generic_eq p0_ "sedeleg" then returnm true
+ else if generic_eq p0_ "sideleg" then returnm true
+ else if generic_eq p0_ "sie" then returnm true
+ else if generic_eq p0_ "stvec" then returnm true
+ else if generic_eq p0_ "scounteren" then returnm true
+ else if generic_eq p0_ "sscratch" then returnm true
+ else if generic_eq p0_ "sepc" then returnm true
+ else if generic_eq p0_ "scause" then returnm true
+ else if generic_eq p0_ "stval" then returnm true
+ else if generic_eq p0_ "sip" then returnm true
+ else if generic_eq p0_ "satp" then returnm true
+ else if generic_eq p0_ "mvendorid" then returnm true
+ else if generic_eq p0_ "marchid" then returnm true
+ else if generic_eq p0_ "mimpid" then returnm true
+ else if generic_eq p0_ "mhartid" then returnm true
+ else if generic_eq p0_ "mstatus" then returnm true
+ else if generic_eq p0_ "misa" then returnm true
+ else if generic_eq p0_ "medeleg" then returnm true
+ else if generic_eq p0_ "mideleg" then returnm true
+ else if generic_eq p0_ "mie" then returnm true
+ else if generic_eq p0_ "mtvec" then returnm true
+ else if generic_eq p0_ "mcounteren" then returnm true
+ else if generic_eq p0_ "mscratch" then returnm true
+ else if generic_eq p0_ "mepc" then returnm true
+ else if generic_eq p0_ "mcause" then returnm true
+ else if generic_eq p0_ "mtval" then returnm true
+ else if generic_eq p0_ "mip" then returnm true
+ else if generic_eq p0_ "pmpcfg0" then returnm true
+ else if generic_eq p0_ "pmpcfg1" then returnm true
+ else if generic_eq p0_ "pmpcfg2" then returnm true
+ else if generic_eq p0_ "pmpcfg3" then returnm true
+ else if generic_eq p0_ "pmpaddr0" then returnm true
+ else if generic_eq p0_ "pmpaddr1" then returnm true
+ else if generic_eq p0_ "pmpaddr2" then returnm true
+ else if generic_eq p0_ "pmpaddr3" then returnm true
+ else if generic_eq p0_ "pmpaddr4" then returnm true
+ else if generic_eq p0_ "pmpaddr5" then returnm true
+ else if generic_eq p0_ "pmpaddr6" then returnm true
+ else if generic_eq p0_ "pmpaddr7" then returnm true
+ else if generic_eq p0_ "pmpaddr8" then returnm true
+ else if generic_eq p0_ "pmpaddr9" then returnm true
+ else if generic_eq p0_ "pmpaddr10" then returnm true
+ else if generic_eq p0_ "pmpaddr11" then returnm true
+ else if generic_eq p0_ "pmpaddr12" then returnm true
+ else if generic_eq p0_ "pmpaddr13" then returnm true
+ else if generic_eq p0_ "pmpaddr14" then returnm true
+ else if generic_eq p0_ "pmpaddr15" then returnm true
+ else if generic_eq p0_ "mcycle" then returnm true
+ else if generic_eq p0_ "minstret" then returnm true
+ else if generic_eq p0_ "mcycleh" then returnm true
+ else if generic_eq p0_ "minstreth" then returnm true
+ else if generic_eq p0_ "tselect" then returnm true
+ else if generic_eq p0_ "tdata1" then returnm true
+ else if generic_eq p0_ "tdata2" then returnm true
+ else if generic_eq p0_ "tdata3" then returnm true
+ else
+ (and_boolM (returnm ((hex_bits_12_backwards_matches p0_) : bool))
+ ((if hex_bits_12_backwards_matches p0_ then
+ (hex_bits_12_backwards p0_) >>= fun reg => returnm (true : bool)
+ else returnm false)
+ : M (bool))) >>= fun w__1 : bool =>
+ if sumbool_of_bool w__1 then (hex_bits_12_backwards p0_) >>= fun reg => returnm true
+ else returnm false)
+ : M (bool).
-Definition lift_uie (o : Sinterrupts) (d : Sinterrupts) (u : Uinterrupts)
-: Sinterrupts :=
-
- let s : Sinterrupts := o in
- let s :=
- if ((eq_vec (_get_Sinterrupts_UEI d) ((bool_to_bits true) : mword 1))) then
- _update_Sinterrupts_UEI s (_get_Uinterrupts_UEI u)
- else s in
- let s :=
- if ((eq_vec (_get_Sinterrupts_UTI d) ((bool_to_bits true) : mword 1))) then
- _update_Sinterrupts_UTI s (_get_Uinterrupts_UTI u)
- else s in
- if ((eq_vec (_get_Sinterrupts_USI d) ((bool_to_bits true) : mword 1))) then
- _update_Sinterrupts_USI s (_get_Uinterrupts_USI u)
- else s.
+Definition _s748_ (_s749_ : string) : option ((mword 12 * string)) :=
+ match _s749_ with
+ | _s750_ =>
+ match (hex_bits_12_matches_prefix _s750_) with
+ | Some (reg, existT _ _s751_ _) =>
+ match (string_drop _s750_ _s751_) with | s_ => Some (reg, s_) end
+ | _ => None
+ end
+ end.
-Definition legalize_uie (s : Sinterrupts) (d : Sinterrupts) (v : mword 64)
-: Sinterrupts :=
-
- lift_uie s d (Mk_Uinterrupts v).
+Definition _s744_ (_s745_ : string) : option string :=
+ let _s746_ := _s745_ in
+ if string_startswith _s746_ "tdata3" then
+ match (string_drop _s746_ (projT1 (string_length "tdata3"))) with | s_ => Some s_ end
+ else None.
-Definition handle_trap_extension (p : Privilege) (pc : mword 64) (u : option unit) : unit := tt.
+Definition _s740_ (_s741_ : string) : option string :=
+ let _s742_ := _s741_ in
+ if string_startswith _s742_ "tdata2" then
+ match (string_drop _s742_ (projT1 (string_length "tdata2"))) with | s_ => Some s_ end
+ else None.
-Definition prepare_trap_vector (p : Privilege) (cause : Mcause)
-: M (mword 64) :=
-
- (match p with
- | Machine => read_reg mtvec_ref : M (Mtvec)
- | Supervisor => read_reg stvec_ref : M (Mtvec)
- | User => read_reg utvec_ref : M (Mtvec)
- end) >>= fun tvec : Mtvec =>
- (match (tvec_addr tvec cause) with
- | Some (epc) => returnm (epc : mword 64)
- | None => (internal_error "Invalid tvec mode") : M (mword 64)
- end)
- : M (mword 64).
+Definition _s736_ (_s737_ : string) : option string :=
+ let _s738_ := _s737_ in
+ if string_startswith _s738_ "tdata1" then
+ match (string_drop _s738_ (projT1 (string_length "tdata1"))) with | s_ => Some s_ end
+ else None.
-Definition get_xret_target (p : Privilege)
-: M (mword 64) :=
-
- (match p with
- | Machine => ((read_reg mepc_ref) : M (mword 64)) : M (mword 64)
- | Supervisor => ((read_reg sepc_ref) : M (mword 64)) : M (mword 64)
- | User => ((read_reg uepc_ref) : M (mword 64)) : M (mword 64)
- end)
- : M (mword 64).
+Definition _s732_ (_s733_ : string) : option string :=
+ let _s734_ := _s733_ in
+ if string_startswith _s734_ "tselect" then
+ match (string_drop _s734_ (projT1 (string_length "tselect"))) with | s_ => Some s_ end
+ else None.
-Definition set_xret_target (p : Privilege) (value : mword 64)
-: M (mword 64) :=
-
- (legalize_xepc value) >>= fun target =>
- (match p with
- | Machine => write_reg mepc_ref target : M (unit)
- | Supervisor => write_reg sepc_ref target : M (unit)
- | User => write_reg uepc_ref target : M (unit)
- end) >>
- returnm (target
- : mword 64).
+Definition _s728_ (_s729_ : string) : option string :=
+ let _s730_ := _s729_ in
+ if string_startswith _s730_ "minstreth" then
+ match (string_drop _s730_ (projT1 (string_length "minstreth"))) with | s_ => Some s_ end
+ else None.
-Definition prepare_xret_target (p : Privilege)
-: M (mword 64) :=
-
- (get_xret_target p)
- : M (mword 64).
+Definition _s724_ (_s725_ : string) : option string :=
+ let _s726_ := _s725_ in
+ if string_startswith _s726_ "mcycleh" then
+ match (string_drop _s726_ (projT1 (string_length "mcycleh"))) with | s_ => Some s_ end
+ else None.
-Definition get_mtvec '(tt : unit)
-: M (mword 64) :=
-
- read_reg mtvec_ref >>= fun w__0 : Mtvec => returnm ((_get_Mtvec_bits w__0) : mword 64).
+Definition _s720_ (_s721_ : string) : option string :=
+ let _s722_ := _s721_ in
+ if string_startswith _s722_ "minstret" then
+ match (string_drop _s722_ (projT1 (string_length "minstret"))) with | s_ => Some s_ end
+ else None.
-Definition get_stvec '(tt : unit)
-: M (mword 64) :=
-
- read_reg stvec_ref >>= fun w__0 : Mtvec => returnm ((_get_Mtvec_bits w__0) : mword 64).
+Definition _s716_ (_s717_ : string) : option string :=
+ let _s718_ := _s717_ in
+ if string_startswith _s718_ "mcycle" then
+ match (string_drop _s718_ (projT1 (string_length "mcycle"))) with | s_ => Some s_ end
+ else None.
-Definition get_utvec '(tt : unit)
-: M (mword 64) :=
-
- read_reg utvec_ref >>= fun w__0 : Mtvec => returnm ((_get_Mtvec_bits w__0) : mword 64).
+Definition _s712_ (_s713_ : string) : option string :=
+ let _s714_ := _s713_ in
+ if string_startswith _s714_ "pmpaddr15" then
+ match (string_drop _s714_ (projT1 (string_length "pmpaddr15"))) with | s_ => Some s_ end
+ else None.
-Definition set_mtvec (value : mword 64)
-: M (mword 64) :=
-
- read_reg mtvec_ref >>= fun w__0 : Mtvec =>
- write_reg mtvec_ref (legalize_tvec w__0 value) >>
- read_reg mtvec_ref >>= fun w__1 : Mtvec => returnm ((_get_Mtvec_bits w__1) : mword 64).
+Definition _s708_ (_s709_ : string) : option string :=
+ let _s710_ := _s709_ in
+ if string_startswith _s710_ "pmpaddr14" then
+ match (string_drop _s710_ (projT1 (string_length "pmpaddr14"))) with | s_ => Some s_ end
+ else None.
-Definition set_stvec (value : mword 64)
-: M (mword 64) :=
-
- read_reg stvec_ref >>= fun w__0 : Mtvec =>
- write_reg stvec_ref (legalize_tvec w__0 value) >>
- read_reg stvec_ref >>= fun w__1 : Mtvec => returnm ((_get_Mtvec_bits w__1) : mword 64).
+Definition _s704_ (_s705_ : string) : option string :=
+ let _s706_ := _s705_ in
+ if string_startswith _s706_ "pmpaddr13" then
+ match (string_drop _s706_ (projT1 (string_length "pmpaddr13"))) with | s_ => Some s_ end
+ else None.
-Definition set_utvec (value : mword 64)
-: M (mword 64) :=
-
- read_reg utvec_ref >>= fun w__0 : Mtvec =>
- write_reg utvec_ref (legalize_tvec w__0 value) >>
- read_reg utvec_ref >>= fun w__1 : Mtvec => returnm ((_get_Mtvec_bits w__1) : mword 64).
+Definition _s700_ (_s701_ : string) : option string :=
+ let _s702_ := _s701_ in
+ if string_startswith _s702_ "pmpaddr12" then
+ match (string_drop _s702_ (projT1 (string_length "pmpaddr12"))) with | s_ => Some s_ end
+ else None.
-Definition is_NExt_CSR_defined (csr : mword 12) (p : Privilege)
-: M (bool) :=
-
- let b__0 := csr in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- (haveUsrMode tt)
+Definition _s696_ (_s697_ : string) : option string :=
+ let _s698_ := _s697_ in
+ if string_startswith _s698_ "pmpaddr11" then
+ match (string_drop _s698_ (projT1 (string_length "pmpaddr11"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s692_ (_s693_ : string) : option string :=
+ let _s694_ := _s693_ in
+ if string_startswith _s694_ "pmpaddr10" then
+ match (string_drop _s694_ (projT1 (string_length "pmpaddr10"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s688_ (_s689_ : string) : option string :=
+ let _s690_ := _s689_ in
+ if string_startswith _s690_ "pmpaddr9" then
+ match (string_drop _s690_ (projT1 (string_length "pmpaddr9"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s684_ (_s685_ : string) : option string :=
+ let _s686_ := _s685_ in
+ if string_startswith _s686_ "pmpaddr8" then
+ match (string_drop _s686_ (projT1 (string_length "pmpaddr8"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s680_ (_s681_ : string) : option string :=
+ let _s682_ := _s681_ in
+ if string_startswith _s682_ "pmpaddr7" then
+ match (string_drop _s682_ (projT1 (string_length "pmpaddr7"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s676_ (_s677_ : string) : option string :=
+ let _s678_ := _s677_ in
+ if string_startswith _s678_ "pmpaddr6" then
+ match (string_drop _s678_ (projT1 (string_length "pmpaddr6"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s672_ (_s673_ : string) : option string :=
+ let _s674_ := _s673_ in
+ if string_startswith _s674_ "pmpaddr5" then
+ match (string_drop _s674_ (projT1 (string_length "pmpaddr5"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s668_ (_s669_ : string) : option string :=
+ let _s670_ := _s669_ in
+ if string_startswith _s670_ "pmpaddr4" then
+ match (string_drop _s670_ (projT1 (string_length "pmpaddr4"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s664_ (_s665_ : string) : option string :=
+ let _s666_ := _s665_ in
+ if string_startswith _s666_ "pmpaddr3" then
+ match (string_drop _s666_ (projT1 (string_length "pmpaddr3"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s660_ (_s661_ : string) : option string :=
+ let _s662_ := _s661_ in
+ if string_startswith _s662_ "pmpaddr2" then
+ match (string_drop _s662_ (projT1 (string_length "pmpaddr2"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s656_ (_s657_ : string) : option string :=
+ let _s658_ := _s657_ in
+ if string_startswith _s658_ "pmpaddr1" then
+ match (string_drop _s658_ (projT1 (string_length "pmpaddr1"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s652_ (_s653_ : string) : option string :=
+ let _s654_ := _s653_ in
+ if string_startswith _s654_ "pmpaddr0" then
+ match (string_drop _s654_ (projT1 (string_length "pmpaddr0"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s648_ (_s649_ : string) : option string :=
+ let _s650_ := _s649_ in
+ if string_startswith _s650_ "pmpcfg3" then
+ match (string_drop _s650_ (projT1 (string_length "pmpcfg3"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s644_ (_s645_ : string) : option string :=
+ let _s646_ := _s645_ in
+ if string_startswith _s646_ "pmpcfg2" then
+ match (string_drop _s646_ (projT1 (string_length "pmpcfg2"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s640_ (_s641_ : string) : option string :=
+ let _s642_ := _s641_ in
+ if string_startswith _s642_ "pmpcfg1" then
+ match (string_drop _s642_ (projT1 (string_length "pmpcfg1"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s636_ (_s637_ : string) : option string :=
+ let _s638_ := _s637_ in
+ if string_startswith _s638_ "pmpcfg0" then
+ match (string_drop _s638_ (projT1 (string_length "pmpcfg0"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s632_ (_s633_ : string) : option string :=
+ let _s634_ := _s633_ in
+ if string_startswith _s634_ "mip" then
+ match (string_drop _s634_ (projT1 (string_length "mip"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s628_ (_s629_ : string) : option string :=
+ let _s630_ := _s629_ in
+ if string_startswith _s630_ "mtval" then
+ match (string_drop _s630_ (projT1 (string_length "mtval"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s624_ (_s625_ : string) : option string :=
+ let _s626_ := _s625_ in
+ if string_startswith _s626_ "mcause" then
+ match (string_drop _s626_ (projT1 (string_length "mcause"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s620_ (_s621_ : string) : option string :=
+ let _s622_ := _s621_ in
+ if string_startswith _s622_ "mepc" then
+ match (string_drop _s622_ (projT1 (string_length "mepc"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s616_ (_s617_ : string) : option string :=
+ let _s618_ := _s617_ in
+ if string_startswith _s618_ "mscratch" then
+ match (string_drop _s618_ (projT1 (string_length "mscratch"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s612_ (_s613_ : string) : option string :=
+ let _s614_ := _s613_ in
+ if string_startswith _s614_ "mcounteren" then
+ match (string_drop _s614_ (projT1 (string_length "mcounteren"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s608_ (_s609_ : string) : option string :=
+ let _s610_ := _s609_ in
+ if string_startswith _s610_ "mtvec" then
+ match (string_drop _s610_ (projT1 (string_length "mtvec"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s604_ (_s605_ : string) : option string :=
+ let _s606_ := _s605_ in
+ if string_startswith _s606_ "mie" then
+ match (string_drop _s606_ (projT1 (string_length "mie"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s600_ (_s601_ : string) : option string :=
+ let _s602_ := _s601_ in
+ if string_startswith _s602_ "mideleg" then
+ match (string_drop _s602_ (projT1 (string_length "mideleg"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s596_ (_s597_ : string) : option string :=
+ let _s598_ := _s597_ in
+ if string_startswith _s598_ "medeleg" then
+ match (string_drop _s598_ (projT1 (string_length "medeleg"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s592_ (_s593_ : string) : option string :=
+ let _s594_ := _s593_ in
+ if string_startswith _s594_ "misa" then
+ match (string_drop _s594_ (projT1 (string_length "misa"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s588_ (_s589_ : string) : option string :=
+ let _s590_ := _s589_ in
+ if string_startswith _s590_ "mstatus" then
+ match (string_drop _s590_ (projT1 (string_length "mstatus"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s584_ (_s585_ : string) : option string :=
+ let _s586_ := _s585_ in
+ if string_startswith _s586_ "mhartid" then
+ match (string_drop _s586_ (projT1 (string_length "mhartid"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s580_ (_s581_ : string) : option string :=
+ let _s582_ := _s581_ in
+ if string_startswith _s582_ "mimpid" then
+ match (string_drop _s582_ (projT1 (string_length "mimpid"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s576_ (_s577_ : string) : option string :=
+ let _s578_ := _s577_ in
+ if string_startswith _s578_ "marchid" then
+ match (string_drop _s578_ (projT1 (string_length "marchid"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s572_ (_s573_ : string) : option string :=
+ let _s574_ := _s573_ in
+ if string_startswith _s574_ "mvendorid" then
+ match (string_drop _s574_ (projT1 (string_length "mvendorid"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s568_ (_s569_ : string) : option string :=
+ let _s570_ := _s569_ in
+ if string_startswith _s570_ "satp" then
+ match (string_drop _s570_ (projT1 (string_length "satp"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s564_ (_s565_ : string) : option string :=
+ let _s566_ := _s565_ in
+ if string_startswith _s566_ "sip" then
+ match (string_drop _s566_ (projT1 (string_length "sip"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s560_ (_s561_ : string) : option string :=
+ let _s562_ := _s561_ in
+ if string_startswith _s562_ "stval" then
+ match (string_drop _s562_ (projT1 (string_length "stval"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s556_ (_s557_ : string) : option string :=
+ let _s558_ := _s557_ in
+ if string_startswith _s558_ "scause" then
+ match (string_drop _s558_ (projT1 (string_length "scause"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s552_ (_s553_ : string) : option string :=
+ let _s554_ := _s553_ in
+ if string_startswith _s554_ "sepc" then
+ match (string_drop _s554_ (projT1 (string_length "sepc"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s548_ (_s549_ : string) : option string :=
+ let _s550_ := _s549_ in
+ if string_startswith _s550_ "sscratch" then
+ match (string_drop _s550_ (projT1 (string_length "sscratch"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s544_ (_s545_ : string) : option string :=
+ let _s546_ := _s545_ in
+ if string_startswith _s546_ "scounteren" then
+ match (string_drop _s546_ (projT1 (string_length "scounteren"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s540_ (_s541_ : string) : option string :=
+ let _s542_ := _s541_ in
+ if string_startswith _s542_ "stvec" then
+ match (string_drop _s542_ (projT1 (string_length "stvec"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s536_ (_s537_ : string) : option string :=
+ let _s538_ := _s537_ in
+ if string_startswith _s538_ "sie" then
+ match (string_drop _s538_ (projT1 (string_length "sie"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s532_ (_s533_ : string) : option string :=
+ let _s534_ := _s533_ in
+ if string_startswith _s534_ "sideleg" then
+ match (string_drop _s534_ (projT1 (string_length "sideleg"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s528_ (_s529_ : string) : option string :=
+ let _s530_ := _s529_ in
+ if string_startswith _s530_ "sedeleg" then
+ match (string_drop _s530_ (projT1 (string_length "sedeleg"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s524_ (_s525_ : string) : option string :=
+ let _s526_ := _s525_ in
+ if string_startswith _s526_ "sstatus" then
+ match (string_drop _s526_ (projT1 (string_length "sstatus"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s520_ (_s521_ : string) : option string :=
+ let _s522_ := _s521_ in
+ if string_startswith _s522_ "instreth" then
+ match (string_drop _s522_ (projT1 (string_length "instreth"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s516_ (_s517_ : string) : option string :=
+ let _s518_ := _s517_ in
+ if string_startswith _s518_ "timeh" then
+ match (string_drop _s518_ (projT1 (string_length "timeh"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s512_ (_s513_ : string) : option string :=
+ let _s514_ := _s513_ in
+ if string_startswith _s514_ "cycleh" then
+ match (string_drop _s514_ (projT1 (string_length "cycleh"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s508_ (_s509_ : string) : option string :=
+ let _s510_ := _s509_ in
+ if string_startswith _s510_ "instret" then
+ match (string_drop _s510_ (projT1 (string_length "instret"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s504_ (_s505_ : string) : option string :=
+ let _s506_ := _s505_ in
+ if string_startswith _s506_ "time" then
+ match (string_drop _s506_ (projT1 (string_length "time"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s500_ (_s501_ : string) : option string :=
+ let _s502_ := _s501_ in
+ if string_startswith _s502_ "cycle" then
+ match (string_drop _s502_ (projT1 (string_length "cycle"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s496_ (_s497_ : string) : option string :=
+ let _s498_ := _s497_ in
+ if string_startswith _s498_ "fcsr" then
+ match (string_drop _s498_ (projT1 (string_length "fcsr"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s492_ (_s493_ : string) : option string :=
+ let _s494_ := _s493_ in
+ if string_startswith _s494_ "frm" then
+ match (string_drop _s494_ (projT1 (string_length "frm"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s488_ (_s489_ : string) : option string :=
+ let _s490_ := _s489_ in
+ if string_startswith _s490_ "fflags" then
+ match (string_drop _s490_ (projT1 (string_length "fflags"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s484_ (_s485_ : string) : option string :=
+ let _s486_ := _s485_ in
+ if string_startswith _s486_ "uip" then
+ match (string_drop _s486_ (projT1 (string_length "uip"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s480_ (_s481_ : string) : option string :=
+ let _s482_ := _s481_ in
+ if string_startswith _s482_ "utval" then
+ match (string_drop _s482_ (projT1 (string_length "utval"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s476_ (_s477_ : string) : option string :=
+ let _s478_ := _s477_ in
+ if string_startswith _s478_ "ucause" then
+ match (string_drop _s478_ (projT1 (string_length "ucause"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s472_ (_s473_ : string) : option string :=
+ let _s474_ := _s473_ in
+ if string_startswith _s474_ "uepc" then
+ match (string_drop _s474_ (projT1 (string_length "uepc"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s468_ (_s469_ : string) : option string :=
+ let _s470_ := _s469_ in
+ if string_startswith _s470_ "uscratch" then
+ match (string_drop _s470_ (projT1 (string_length "uscratch"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s464_ (_s465_ : string) : option string :=
+ let _s466_ := _s465_ in
+ if string_startswith _s466_ "utvec" then
+ match (string_drop _s466_ (projT1 (string_length "utvec"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s460_ (_s461_ : string) : option string :=
+ let _s462_ := _s461_ in
+ if string_startswith _s462_ "uie" then
+ match (string_drop _s462_ (projT1 (string_length "uie"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s456_ (_s457_ : string) : option string :=
+ let _s458_ := _s457_ in
+ if string_startswith _s458_ "ustatus" then
+ match (string_drop _s458_ (projT1 (string_length "ustatus"))) with | s_ => Some s_ end
+ else None.
+
+Definition csr_name_map_matches_prefix (arg_ : string)
+: M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s459_ := arg_ in
+ (if match (_s456_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s456_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"000"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s460_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s460_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"004"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s464_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s464_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"005"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s468_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s468_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"040"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s472_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s472_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"041"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s476_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s476_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"042"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s480_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s480_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"043"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s484_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s484_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"044"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s488_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s488_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"001"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s492_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s492_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"002"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s496_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s496_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"003"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s500_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s500_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"C00"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s504_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s504_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"C01"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s508_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s508_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"C02"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s512_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s512_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"C80"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s516_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s516_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"C81"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s520_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s520_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"C82"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s524_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s524_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"100"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s528_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s528_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"102"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s532_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s532_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"103"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s536_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s536_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"104"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s540_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s540_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"105"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s544_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s544_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"106"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s548_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s548_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"140"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s552_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s552_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"141"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s556_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s556_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"142"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s560_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s560_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"143"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s564_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s564_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"144"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s568_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s568_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"180"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s572_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s572_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"F11"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s576_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s576_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"F12"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s580_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s580_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"F13"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s584_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s584_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"F14"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s588_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s588_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"300"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s592_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s592_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"301"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s596_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s596_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"302"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s600_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s600_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"303"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s604_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s604_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"304"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s608_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s608_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"305"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s612_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s612_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"306"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s616_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s616_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"340"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s620_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s620_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"341"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s624_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s624_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"342"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s628_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s628_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"343"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s632_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s632_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"344"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s636_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s636_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3A0"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s640_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s640_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3A1"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s644_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s644_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3A2"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s648_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s648_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3A3"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s652_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s652_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B0"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s656_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s656_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B1"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s660_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s660_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B2"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s664_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s664_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B3"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s668_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s668_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B4"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s672_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s672_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B5"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s676_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s676_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B6"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s680_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s680_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B7"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s684_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s684_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B8"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s688_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s688_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3B9"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s692_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s692_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3BA"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s696_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s696_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3BB"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s700_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s700_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3BC"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s704_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s704_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3BD"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s708_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s708_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3BE"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s712_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s712_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"3BF"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s716_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s716_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"B00"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s720_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s720_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"B02"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s724_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s724_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"B80"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s728_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s728_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"B82"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s732_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s732_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"7A0"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s736_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s736_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"7A1"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s740_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s740_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"7A2"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s744_ _s459_) with | Some s_ => true | _ => false end then
+ (match (_s744_ _s459_) with
+ | Some s_ =>
+ returnm (Some
+ (Ox"7A3"
+ : mword 12, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s748_ _s459_) with | Some (reg, s_) => true | _ => false end then
+ (match (_s748_ _s459_) with
+ | Some (reg, s_) =>
+ returnm (Some
+ (reg, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 12 * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition csr_name (csr : mword 12) : string := csr_name_map_forwards csr.
+
+Definition ext_is_CSR_defined (b__0 : mword 12) (g__325 : Privilege) : M (bool) :=
+ (if eq_vec b__0 (Ox"000" : mword 12) then
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool)))
+ : M (bool)
+ else if eq_vec b__0 (Ox"004" : mword 12) then
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool)))
+ : M (bool)
+ else if eq_vec b__0 (Ox"005" : mword 12) then
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool)))
+ : M (bool)
+ else if eq_vec b__0 (Ox"040" : mword 12) then
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- (haveUsrMode tt)
+ else if eq_vec b__0 (Ox"041" : mword 12) then
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- (haveUsrMode tt)
+ else if eq_vec b__0 (Ox"042" : mword 12) then
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- (haveUsrMode tt)
+ else if eq_vec b__0 (Ox"043" : mword 12) then
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- (haveUsrMode tt)
+ else if eq_vec b__0 (Ox"044" : mword 12) then
+ (and_boolM ((haveUsrMode tt) : M (bool)) ((haveNExt tt) : M (bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- (haveUsrMode tt)
+ else if eq_vec b__0 (Ox"001" : mword 12) then
+ (or_boolM ((haveFExt tt) : M (bool)) ((haveDExt tt) : M (bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- (haveUsrMode tt)
+ else if eq_vec b__0 (Ox"002" : mword 12) then
+ (or_boolM ((haveFExt tt) : M (bool)) ((haveDExt tt) : M (bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- (haveUsrMode tt)
+ else if eq_vec b__0 (Ox"003" : mword 12) then
+ (or_boolM ((haveFExt tt) : M (bool)) ((haveDExt tt) : M (bool)))
: M (bool)
- else returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))})))
+ else returnm false)
: M (bool).
-Definition read_NExt_CSR (csr : mword 12)
-: M (option (mword 64)) :=
-
- let b__0 := csr in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+Definition ext_read_CSR (b__0 : mword 12) : M (option (mword 64)) :=
+ (if eq_vec b__0 (Ox"000" : mword 12) then
read_reg mstatus_ref >>= fun w__0 : Mstatus =>
- returnm ((Some
- (_get_Ustatus_bits (lower_sstatus (lower_mstatus w__0))))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ returnm (Some (_get_Ustatus_bits (lower_sstatus (lower_mstatus w__0))))
+ else if eq_vec b__0 (Ox"004" : mword 12) then
read_reg mie_ref >>= fun w__1 : Minterrupts =>
read_reg mideleg_ref >>= fun w__2 : Minterrupts =>
read_reg sideleg_ref >>= fun w__3 : Sinterrupts =>
- returnm ((Some
- (_get_Uinterrupts_bits (lower_sie (lower_mie w__1 w__2) w__3)))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- (get_utvec tt) >>= fun w__4 : mword 64 => returnm ((Some (w__4)) : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- ((read_reg uscratch_ref) : M (mword 64)) >>= fun w__5 : mword 64 =>
- returnm ((Some
- (w__5))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ returnm (Some (_get_Uinterrupts_bits (lower_sie (lower_mie w__1 w__2) w__3)))
+ else if eq_vec b__0 (Ox"005" : mword 12) then
+ (get_utvec tt) >>= fun w__4 : mword 64 => returnm (Some w__4)
+ else if eq_vec b__0 (Ox"040" : mword 12) then
+ ((read_reg uscratch_ref) : M (mword 64)) >>= fun w__5 : mword 64 => returnm (Some w__5)
+ else if eq_vec b__0 (Ox"041" : mword 12) then
(get_xret_target User) >>= fun w__6 : mword 64 =>
- (pc_alignment_mask tt) >>= fun w__7 : mword 64 =>
- returnm ((Some
- (and_vec w__6 w__7))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg ucause_ref >>= fun w__8 : Mcause =>
- returnm ((Some
- (_get_Mcause_bits w__8))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- ((read_reg utval_ref) : M (mword 64)) >>= fun w__9 : mword 64 =>
- returnm ((Some
- (w__9))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ (pc_alignment_mask tt) >>= fun w__7 : mword 64 => returnm (Some (and_vec w__6 w__7))
+ else if eq_vec b__0 (Ox"042" : mword 12) then
+ read_reg ucause_ref >>= fun w__8 : Mcause => returnm (Some (_get_Mcause_bits w__8))
+ else if eq_vec b__0 (Ox"043" : mword 12) then
+ ((read_reg utval_ref) : M (mword 64)) >>= fun w__9 : mword 64 => returnm (Some w__9)
+ else if eq_vec b__0 (Ox"044" : mword 12) then
read_reg mip_ref >>= fun w__10 : Minterrupts =>
read_reg mideleg_ref >>= fun w__11 : Minterrupts =>
read_reg sideleg_ref >>= fun w__12 : Sinterrupts =>
- returnm ((Some
- (_get_Uinterrupts_bits (lower_sip (lower_mip w__10 w__11) w__12)))
- : option (mword 64))
- else returnm (None : option (mword 64)))
- : M (option xlenbits).
-
-Definition write_NExt_CSR (csr : mword 12) (value : mword 64)
-: M (bool) :=
-
- let b__0 := csr in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ returnm (Some (_get_Uinterrupts_bits (lower_sip (lower_mip w__10 w__11) w__12)))
+ else if eq_vec b__0 (Ox"001" : mword 12) then
+ read_reg fcsr_ref >>= fun w__13 : Fcsr => returnm (Some (EXTZ 64 (_get_Fcsr_FFLAGS w__13)))
+ else if eq_vec b__0 (Ox"002" : mword 12) then
+ read_reg fcsr_ref >>= fun w__14 : Fcsr => returnm (Some (EXTZ 64 (_get_Fcsr_FRM w__14)))
+ else if eq_vec b__0 (Ox"003" : mword 12) then
+ read_reg fcsr_ref >>= fun w__15 : Fcsr => returnm (Some (EXTZ 64 (_get_Fcsr_bits w__15)))
+ else returnm None)
+ : M (option (mword 64)).
+
+Definition ext_write_CSR (b__0 : mword 12) (value : mword 64) : M (option (mword 64)) :=
+ (if eq_vec b__0 (Ox"000" : mword 12) then
read_reg mstatus_ref >>= fun w__0 : Mstatus =>
(legalize_ustatus w__0 value) >>= fun w__1 : Mstatus =>
write_reg mstatus_ref w__1 >>
- read_reg mstatus_ref >>= fun w__2 : Mstatus =>
- returnm ((Some
- (_get_Mstatus_bits w__2))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ read_reg mstatus_ref >>= fun w__2 : Mstatus => returnm (Some (_get_Mstatus_bits w__2))
+ else if eq_vec b__0 (Ox"004" : mword 12) then
read_reg mie_ref >>= fun w__3 : Minterrupts =>
read_reg mideleg_ref >>= fun w__4 : Minterrupts =>
read_reg sideleg_ref >>= fun w__5 : Sinterrupts =>
@@ -10319,36 +10494,21 @@ Definition write_NExt_CSR (csr : mword 12) (value : mword 64)
read_reg mideleg_ref >>= fun w__7 : Minterrupts =>
(lift_sie w__6 w__7 sie) >>= fun w__8 : Minterrupts =>
write_reg mie_ref w__8 >>
- read_reg mie_ref >>= fun w__9 : Minterrupts =>
- returnm ((Some
- (_get_Minterrupts_bits w__9))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- (set_utvec value) >>= fun w__10 : mword 64 => returnm ((Some (w__10)) : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ read_reg mie_ref >>= fun w__9 : Minterrupts => returnm (Some (_get_Minterrupts_bits w__9))
+ else if eq_vec b__0 (Ox"005" : mword 12) then
+ (set_utvec value) >>= fun w__10 : mword 64 => returnm (Some w__10)
+ else if eq_vec b__0 (Ox"040" : mword 12) then
write_reg uscratch_ref value >>
- ((read_reg uscratch_ref) : M (mword 64)) >>= fun w__11 : mword 64 =>
- returnm ((Some
- (w__11))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- (set_xret_target User value) >>= fun w__12 : mword 64 =>
- returnm ((Some
- (w__12))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ ((read_reg uscratch_ref) : M (mword 64)) >>= fun w__11 : mword 64 => returnm (Some w__11)
+ else if eq_vec b__0 (Ox"041" : mword 12) then
+ (set_xret_target User value) >>= fun w__12 : mword 64 => returnm (Some w__12)
+ else if eq_vec b__0 (Ox"042" : mword 12) then
(_set_Mcause_bits ucause_ref value) >>
- read_reg ucause_ref >>= fun w__13 : Mcause =>
- returnm ((Some
- (_get_Mcause_bits w__13))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ read_reg ucause_ref >>= fun w__13 : Mcause => returnm (Some (_get_Mcause_bits w__13))
+ else if eq_vec b__0 (Ox"043" : mword 12) then
write_reg utval_ref value >>
- ((read_reg utval_ref) : M (mword 64)) >>= fun w__14 : mword 64 =>
- returnm ((Some
- (w__14))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ ((read_reg utval_ref) : M (mword 64)) >>= fun w__14 : mword 64 => returnm (Some w__14)
+ else if eq_vec b__0 (Ox"044" : mword 12) then
read_reg mip_ref >>= fun w__15 : Minterrupts =>
read_reg mideleg_ref >>= fun w__16 : Minterrupts =>
read_reg sideleg_ref >>= fun w__17 : Sinterrupts =>
@@ -10357,401 +10517,209 @@ Definition write_NExt_CSR (csr : mword 12) (value : mword 64)
read_reg mideleg_ref >>= fun w__19 : Minterrupts =>
(lift_sip w__18 w__19 sip) >>= fun w__20 : Minterrupts =>
write_reg mip_ref w__20 >>
- read_reg mip_ref >>= fun w__21 : Minterrupts =>
- returnm ((Some
- (_get_Minterrupts_bits w__21))
- : option (mword 64))
- else returnm (None : option (mword 64))) >>= fun res : option xlenbits =>
- returnm ((match res with
- | Some (v) =>
- let '_ :=
- (if ((get_config_print_reg tt)) then
- print_endline
- (String.append "CSR "
- (String.append (csr_name csr)
- (String.append " <- "
- (String.append (string_of_bits v)
- (String.append " (input: "
- (String.append (string_of_bits value) ")"))))))
- else tt)
- : unit in
- true
- | None => false
- end)
- : bool).
-
-Definition ext_is_CSR_defined (csr : mword 12) (p : Privilege)
-: M (bool) :=
-
- (is_NExt_CSR_defined csr p)
- : M (bool).
-
-Definition ext_read_CSR (csr : mword 12)
-: M (option (mword 64)) :=
-
- (read_NExt_CSR csr)
+ read_reg mip_ref >>= fun w__21 : Minterrupts => returnm (Some (_get_Minterrupts_bits w__21))
+ else if eq_vec b__0 (Ox"001" : mword 12) then
+ read_reg fcsr_ref >>= fun w__22 : Fcsr =>
+ (ext_write_fcsr (_get_Fcsr_FRM w__22) (subrange_vec_dec value 4 0)) >>
+ read_reg fcsr_ref >>= fun w__23 : Fcsr => returnm (Some (EXTZ 64 (_get_Fcsr_FFLAGS w__23)))
+ else if eq_vec b__0 (Ox"002" : mword 12) then
+ read_reg fcsr_ref >>= fun w__24 : Fcsr =>
+ (ext_write_fcsr (subrange_vec_dec value 2 0) (_get_Fcsr_FFLAGS w__24)) >>
+ read_reg fcsr_ref >>= fun w__25 : Fcsr => returnm (Some (EXTZ 64 (_get_Fcsr_FRM w__25)))
+ else if eq_vec b__0 (Ox"003" : mword 12) then
+ (ext_write_fcsr (subrange_vec_dec value 7 5) (subrange_vec_dec value 4 0)) >>
+ read_reg fcsr_ref >>= fun w__26 : Fcsr => returnm (Some (EXTZ 64 (_get_Fcsr_bits w__26)))
+ else returnm None)
: M (option (mword 64)).
-Definition ext_write_CSR (csr : mword 12) (value : mword 64)
-: M (bool) :=
-
- (write_NExt_CSR csr value)
- : M (bool).
-
-Definition csrAccess (csr : mword 12) : mword 2 := subrange_vec_dec csr 11 10.
+Definition csrAccess (csr : mword 12) : mword 2 := subrange_vec_dec csr 11 10.
-Definition csrPriv (csr : mword 12) : mword 2 := subrange_vec_dec csr 9 8.
+Definition csrPriv (csr : mword 12) : mword 2 := subrange_vec_dec csr 9 8.
-Definition is_CSR_defined (csr : mword 12) (p : Privilege)
-: M (bool) :=
-
+Definition is_CSR_defined (csr : mword 12) (p : Privilege) : M (bool) :=
let b__0 := csr in
- (if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B1] : mword 12))) then
- returnm (projT1
- (build_ex
- (andb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2)) (Z.eqb 64 32))
- : {_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (simp_0 = true /\ 64 = 32))}))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B1] : mword 12))) then
- returnm (projT1
- (build_ex
- (andb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2)) (Z.eqb 64 32))
- : {_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (simp_0 = true /\ 64 = 32))}))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm (projT1
- (build_ex
- (andb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2)) (Z.eqb 64 32))
- : {_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (simp_0 = true /\ 64 = 32))}))
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm (projT1
- (build_ex
- (andb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2)) (Z.eqb 64 32))
- : {_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (simp_0 = true /\ 64 = 32))}))
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ (if eq_vec b__0 (Ox"F11" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"F12" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"F13" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"F14" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"300" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"301" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"302" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"303" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"304" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"305" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"306" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"340" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"341" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"342" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"343" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"344" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3A0" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3A1" : mword 12) then
+ returnm (andb (generic_eq p Machine) (Z.eqb 64 32))
+ else if eq_vec b__0 (Ox"3A2" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3A3" : mword 12) then
+ returnm (andb (generic_eq p Machine) (Z.eqb 64 32))
+ else if eq_vec b__0 (Ox"3B0" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B1" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B2" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B3" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B4" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B5" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B6" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B7" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B8" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3B9" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3BA" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3BB" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3BC" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3BD" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3BE" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"3BF" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"B00" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"B02" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"B80" : mword 12) then
+ returnm (andb (generic_eq p Machine) (Z.eqb 64 32))
+ else if eq_vec b__0 (Ox"B82" : mword 12) then
+ returnm (andb (generic_eq p Machine) (Z.eqb 64 32))
+ else if eq_vec b__0 (Ox"7A0" : mword 12) then returnm (generic_eq p Machine)
+ else if eq_vec b__0 (Ox"100" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"102" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"103" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"104" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"105" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"106" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"140" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"141" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"142" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"143" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"144" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"180" : mword 12) then
(and_boolM ((haveSupMode tt) : M (bool))
- (returnm ((orb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Machine) : mword 2))
- (eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2)))
- : bool)))
+ (returnm ((orb (generic_eq p Machine) (generic_eq p Supervisor)) : bool)))
: M (bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits User) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits User) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits User) : mword 2))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- returnm (projT1
- (build_ex
- (andb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits User) : mword 2)) (Z.eqb 64 32))
- : {_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (simp_0 = true /\ 64 = 32))}))
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- returnm (projT1
- (build_ex
- (andb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits User) : mword 2)) (Z.eqb 64 32))
- : {_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (simp_0 = true /\ 64 = 32))}))
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- returnm (projT1
- (build_ex
- (andb (eq_vec (privLevel_to_bits p) ((privLevel_to_bits User) : mword 2)) (Z.eqb 64 32))
- : {_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (simp_0 = true /\ 64 = 32))}))
+ else if eq_vec b__0 (Ox"C00" : mword 12) then returnm (generic_eq p User)
+ else if eq_vec b__0 (Ox"C01" : mword 12) then returnm (generic_eq p User)
+ else if eq_vec b__0 (Ox"C02" : mword 12) then returnm (generic_eq p User)
+ else if eq_vec b__0 (Ox"C80" : mword 12) then returnm (andb (generic_eq p User) (Z.eqb 64 32))
+ else if eq_vec b__0 (Ox"C81" : mword 12) then returnm (andb (generic_eq p User) (Z.eqb 64 32))
+ else if eq_vec b__0 (Ox"C82" : mword 12) then returnm (andb (generic_eq p User) (Z.eqb 64 32))
else (ext_is_CSR_defined csr p) : M (bool))
: M (bool).
-Definition check_CSR_access (csrrw : mword 2) (csrpr : mword 2) (p : Privilege) (isWrite : bool)
+Definition check_CSR_access (csrrw : mword 2) (csrpr : mword 2) (p : Privilege) (isWrite : bool)
: bool :=
-
- andb (negb (andb (Bool.eqb isWrite true) (eq_vec csrrw (vec_of_bits [B1;B1] : mword 2))))
+ andb (negb (andb (Bool.eqb isWrite true) (eq_vec csrrw ('b"11" : mword 2))))
(zopz0zKzJ_u (privLevel_to_bits p) csrpr).
-Definition check_TVM_SATP (csr : mword 12) (p : Privilege)
-: M (bool) :=
-
- (and_boolM
- (returnm ((eq_vec csr (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))
- : bool))
- ((and_boolM
- (returnm ((eq_vec (privLevel_to_bits p) ((privLevel_to_bits Supervisor) : mword 2))
- : bool))
+Definition check_TVM_SATP (csr : mword 12) (p : Privilege) : M (bool) :=
+ (and_boolM (returnm ((eq_vec csr (Ox"180" : mword 12)) : bool))
+ ((and_boolM (returnm ((generic_eq p Supervisor) : bool))
(read_reg mstatus_ref >>= fun w__0 : Mstatus =>
- returnm ((eq_vec (_get_Mstatus_TVM w__0) ((bool_to_bits true) : mword 1))
- : bool)))
+ returnm ((eq_vec (_get_Mstatus_TVM w__0) ('b"1" : mword 1)) : bool)))
: M (bool))) >>= fun w__2 : bool =>
- returnm ((negb w__2)
- : bool).
+ returnm (negb w__2).
-Definition check_Counteren (csr : mword 12) (p : Privilege)
-: M (bool) :=
-
+Definition check_Counteren (csr : mword 12) (p : Privilege) : M (bool) :=
(match (csr, p) with
| (b__0, Supervisor) =>
- (if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ (if eq_vec b__0 (Ox"C00" : mword 12) then
read_reg mcounteren_ref >>= fun w__0 : Counteren =>
- returnm ((eq_vec (_get_Counteren_CY w__0) ((bool_to_bits true) : mword 1))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ returnm (eq_vec (_get_Counteren_CY w__0) ('b"1" : mword 1))
+ else if eq_vec b__0 (Ox"C01" : mword 12) then
read_reg mcounteren_ref >>= fun w__1 : Counteren =>
- returnm ((eq_vec (_get_Counteren_TM w__1) ((bool_to_bits true) : mword 1))
- : bool)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ returnm (eq_vec (_get_Counteren_TM w__1) ('b"1" : mword 1))
+ else if eq_vec b__0 (Ox"C02" : mword 12) then
read_reg mcounteren_ref >>= fun w__2 : Counteren =>
- returnm ((eq_vec (_get_Counteren_IR w__2) ((bool_to_bits true) : mword 1))
- : bool)
+ returnm (eq_vec (_get_Counteren_IR w__2) ('b"1" : mword 1))
else
- returnm ((match (b__0, Supervisor) with
- | (_, _) =>
- if ((andb
- (zopz0zIzJ_u
- (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12) csr)
- (zopz0zIzJ_u csr
- (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B1;B1;B1;B1;B1] : mword 12))))
- then
- false
- else true
- end)
- : bool))
+ returnm (match (b__0, Supervisor) with
+ | (_, _) =>
+ if andb (zopz0zIzJ_u (Ox"C03" : mword 12) csr)
+ (zopz0zIzJ_u csr (Ox"C1F" : mword 12)) then
+ false
+ else true
+ end))
: M (bool)
| (b__3, User) =>
- (if ((eq_vec b__3 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ (if eq_vec b__3 (Ox"C00" : mword 12) then
(and_boolM
(read_reg mcounteren_ref >>= fun w__6 : Counteren =>
- returnm ((eq_vec (_get_Counteren_CY w__6) ((bool_to_bits true) : mword 1))
- : bool))
+ returnm ((eq_vec (_get_Counteren_CY w__6) ('b"1" : mword 1)) : bool))
((or_boolM ((haveSupMode tt) >>= fun w__7 : bool => returnm ((negb w__7) : bool))
(read_reg scounteren_ref >>= fun w__8 : Counteren =>
- returnm ((eq_vec (_get_Counteren_CY w__8) ((bool_to_bits true) : mword 1))
- : bool)))
+ returnm ((eq_vec (_get_Counteren_CY w__8) ('b"1" : mword 1)) : bool)))
: M (bool)))
: M (bool)
- else if ((eq_vec b__3 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ else if eq_vec b__3 (Ox"C01" : mword 12) then
(and_boolM
(read_reg mcounteren_ref >>= fun w__11 : Counteren =>
- returnm ((eq_vec (_get_Counteren_TM w__11) ((bool_to_bits true) : mword 1))
- : bool))
+ returnm ((eq_vec (_get_Counteren_TM w__11) ('b"1" : mword 1)) : bool))
((or_boolM ((haveSupMode tt) >>= fun w__12 : bool => returnm ((negb w__12) : bool))
(read_reg scounteren_ref >>= fun w__13 : Counteren =>
- returnm ((eq_vec (_get_Counteren_TM w__13) ((bool_to_bits true) : mword 1))
- : bool)))
+ returnm ((eq_vec (_get_Counteren_TM w__13) ('b"1" : mword 1)) : bool)))
: M (bool)))
: M (bool)
- else if ((eq_vec b__3 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ else if eq_vec b__3 (Ox"C02" : mword 12) then
(and_boolM
(read_reg mcounteren_ref >>= fun w__16 : Counteren =>
- returnm ((eq_vec (_get_Counteren_IR w__16) ((bool_to_bits true) : mword 1))
- : bool))
+ returnm ((eq_vec (_get_Counteren_IR w__16) ('b"1" : mword 1)) : bool))
((or_boolM ((haveSupMode tt) >>= fun w__17 : bool => returnm ((negb w__17) : bool))
(read_reg scounteren_ref >>= fun w__18 : Counteren =>
- returnm ((eq_vec (_get_Counteren_IR w__18) ((bool_to_bits true) : mword 1))
- : bool)))
+ returnm ((eq_vec (_get_Counteren_IR w__18) ('b"1" : mword 1)) : bool)))
: M (bool)))
: M (bool)
else
- returnm ((match (b__3, User) with
- | (_, _) =>
- if ((andb
- (zopz0zIzJ_u
- (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12) csr)
- (zopz0zIzJ_u csr
- (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B1;B1;B1;B1;B1] : mword 12))))
- then
- false
- else true
- end)
- : bool))
+ returnm (match (b__3, User) with
+ | (_, _) =>
+ if andb (zopz0zIzJ_u (Ox"C03" : mword 12) csr)
+ (zopz0zIzJ_u csr (Ox"C1F" : mword 12)) then
+ false
+ else true
+ end))
: M (bool)
| (_, _) =>
- returnm ((if ((andb
- (zopz0zIzJ_u (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12)
- csr)
- (zopz0zIzJ_u csr
- (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B1;B1;B1;B1;B1] : mword 12)))) then
- false
- else true)
- : bool)
+ returnm (if andb (zopz0zIzJ_u (Ox"C03" : mword 12) csr)
+ (zopz0zIzJ_u csr (Ox"C1F" : mword 12)) then
+ false
+ else true)
end)
: M (bool).
-Definition check_CSR (csr : mword 12) (p : Privilege) (isWrite : bool)
-: M (bool) :=
-
+Definition check_CSR (csr : mword 12) (p : Privilege) (isWrite : bool) : M (bool) :=
(and_boolM ((is_CSR_defined csr p) : M (bool))
((and_boolM (returnm ((check_CSR_access (csrAccess csr) (csrPriv csr) p isWrite) : bool))
((and_boolM ((check_TVM_SATP csr p) : M (bool)) ((check_Counteren csr p) : M (bool)))
@@ -10767,223 +10735,166 @@ Axiom match_reservation : forall (_ : mword 64) , bool.
Axiom cancel_reservation : forall (_ : unit) , unit.
-Definition exception_delegatee (e : ExceptionType) (p : Privilege)
-: M (Privilege) :=
-
+Definition exception_delegatee (e : ExceptionType) (p : Privilege) : M (Privilege) :=
let idx := projT1 (num_of_ExceptionType e) in
read_reg medeleg_ref >>= fun w__0 : Medeleg =>
- let super := access_vec_dec (_get_Medeleg_bits w__0) idx in
+ (bit_to_bool (access_vec_dec (_get_Medeleg_bits w__0) idx)) >>= fun super =>
(haveSupMode tt) >>= fun w__1 : bool =>
- (if sumbool_of_bool (w__1) then
- (and_boolM ((bit_to_bool super) : M (bool))
- ((and_boolM ((haveNExt tt) : M (bool))
- (read_reg sedeleg_ref >>= fun w__4 : Sedeleg =>
- (bit_to_bool (access_vec_dec (_get_Sedeleg_bits w__4) idx))
- : M (bool)))
- : M (bool)))
- : M (bool)
- else (and_boolM ((bit_to_bool super) : M (bool)) ((haveNExt tt) : M (bool))) : M (bool)) >>= fun user =>
+ (if sumbool_of_bool w__1 then
+ and_boolMP
+ ((returnm (build_ex super)) : M ({_bool : bool & ArithFact (Bool.eqb super _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveNExt tt) : M (bool))
+ (read_reg sedeleg_ref >>= fun w__3 : Sedeleg =>
+ (bit_to_bool (access_vec_dec (_get_Sedeleg_bits w__3) idx))
+ : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (super && simp_0) _bool = true)})
+ else
+ and_boolMP
+ ((returnm (build_ex super)) : M ({_bool : bool & ArithFact (Bool.eqb super _bool)}))
+ (build_trivial_ex
+ ((haveNExt tt)
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (super && simp_0) _bool = true)})) >>= fun '(existT _ user _) =>
(and_boolMP (build_trivial_ex ((haveUsrMode tt) : M (bool)))
- ((returnm (build_ex
- user)) : M ({_bool : bool & ArithFact (iff (_bool = true) (user = true))})) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (user = true /\ simp_0 = true))})) >>= fun '(existT _ w__12 _) =>
- (if sumbool_of_bool (w__12) then returnm (User : Privilege)
+ ((returnm (build_ex user)) : M ({_bool : bool & ArithFact (Bool.eqb user _bool)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (user &&
+ simp_0) _bool = true)})) >>= fun '(existT _ w__10 _) =>
+ (if sumbool_of_bool w__10 then returnm User
else
- (and_boolM ((haveSupMode tt) : M (bool)) ((bit_to_bool super) : M (bool))) >>= fun w__15 : bool =>
- returnm ((if sumbool_of_bool (w__15) then Supervisor
- else Machine)
- : Privilege)) >>= fun deleg =>
- returnm ((if ((zopz0zI_u (privLevel_to_bits deleg) (privLevel_to_bits p))) then p
- else deleg)
- : Privilege).
-
-Definition findPendingInterrupt (ip : mword 64)
-: option InterruptType :=
-
+ (and_boolMP (build_trivial_ex ((haveSupMode tt) : M (bool)))
+ ((returnm (build_ex super)) : M ({_bool : bool & ArithFact (Bool.eqb super _bool)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (super &&
+ simp_0) _bool = true)})) >>= fun '(existT _ w__12 _) =>
+ returnm (if sumbool_of_bool w__12 then Supervisor else Machine)) >>= fun deleg =>
+ returnm (if zopz0zI_u (privLevel_to_bits deleg) (privLevel_to_bits p) then p else deleg).
+
+Definition findPendingInterrupt (ip : mword 64) : option InterruptType :=
let ip := Mk_Minterrupts ip in
- if ((eq_vec (_get_Minterrupts_MEI ip) ((bool_to_bits true) : mword 1))) then Some (I_M_External)
- else if ((eq_vec (_get_Minterrupts_MSI ip) ((bool_to_bits true) : mword 1))) then
- Some
- (I_M_Software)
- else if ((eq_vec (_get_Minterrupts_MTI ip) ((bool_to_bits true) : mword 1))) then
- Some
- (I_M_Timer)
- else if ((eq_vec (_get_Minterrupts_SEI ip) ((bool_to_bits true) : mword 1))) then
- Some
- (I_S_External)
- else if ((eq_vec (_get_Minterrupts_SSI ip) ((bool_to_bits true) : mword 1))) then
- Some
- (I_S_Software)
- else if ((eq_vec (_get_Minterrupts_STI ip) ((bool_to_bits true) : mword 1))) then
- Some
- (I_S_Timer)
- else if ((eq_vec (_get_Minterrupts_UEI ip) ((bool_to_bits true) : mword 1))) then
- Some
- (I_U_External)
- else if ((eq_vec (_get_Minterrupts_USI ip) ((bool_to_bits true) : mword 1))) then
- Some
- (I_U_Software)
- else if ((eq_vec (_get_Minterrupts_UTI ip) ((bool_to_bits true) : mword 1))) then
- Some
- (I_U_Timer)
+ if eq_vec (_get_Minterrupts_MEI ip) ('b"1" : mword 1) then Some I_M_External
+ else if eq_vec (_get_Minterrupts_MSI ip) ('b"1" : mword 1) then Some I_M_Software
+ else if eq_vec (_get_Minterrupts_MTI ip) ('b"1" : mword 1) then Some I_M_Timer
+ else if eq_vec (_get_Minterrupts_SEI ip) ('b"1" : mword 1) then Some I_S_External
+ else if eq_vec (_get_Minterrupts_SSI ip) ('b"1" : mword 1) then Some I_S_Software
+ else if eq_vec (_get_Minterrupts_STI ip) ('b"1" : mword 1) then Some I_S_Timer
+ else if eq_vec (_get_Minterrupts_UEI ip) ('b"1" : mword 1) then Some I_U_External
+ else if eq_vec (_get_Minterrupts_USI ip) ('b"1" : mword 1) then Some I_U_Software
+ else if eq_vec (_get_Minterrupts_UTI ip) ('b"1" : mword 1) then Some I_U_Timer
else None.
Definition processPending
-(xip : Minterrupts) (xie : Minterrupts) (xideleg : mword 64) (priv_enabled : bool)
+(xip : Minterrupts) (xie : Minterrupts) (xideleg : mword 64) (priv_enabled : bool)
: interrupt_set :=
-
let effective_pend :=
and_vec (_get_Minterrupts_bits xip) (and_vec (_get_Minterrupts_bits xie) (not_vec xideleg)) in
let effective_delg := and_vec (_get_Minterrupts_bits xip) xideleg in
- if sumbool_of_bool ((andb priv_enabled
- (neq_vec effective_pend (EXTZ 64 (vec_of_bits [B0] : mword 1))))) then
- Ints_Pending
- (effective_pend)
- else if ((neq_vec effective_delg (EXTZ 64 (vec_of_bits [B0] : mword 1)))) then
- Ints_Delegated
- (effective_delg)
- else Ints_Empty (tt).
-
-Definition getPendingSet (priv : Privilege)
-: M (option ((mword 64 * Privilege))) :=
-
+ if sumbool_of_bool (andb priv_enabled (neq_vec effective_pend (EXTZ 64 ('b"0" : mword 1)))) then
+ Ints_Pending effective_pend
+ else if neq_vec effective_delg (EXTZ 64 ('b"0" : mword 1)) then Ints_Delegated effective_delg
+ else Ints_Empty tt.
+
+Definition getPendingSet (priv : Privilege) : M (option ((mword 64 * Privilege))) :=
(haveUsrMode tt) >>= fun w__0 : bool =>
assert_exp' w__0 "no user mode: M/U or M/S/U system required" >>= fun _ =>
read_reg mip_ref >>= fun w__1 : Minterrupts =>
read_reg mie_ref >>= fun w__2 : Minterrupts =>
let effective_pending := and_vec (_get_Minterrupts_bits w__1) (_get_Minterrupts_bits w__2) in
- (if ((eq_vec effective_pending (EXTZ 64 (vec_of_bits [B0] : mword 1)))) then
- returnm (None
- : option ((mword 64 * Privilege)))
+ (if eq_vec effective_pending (EXTZ 64 ('b"0" : mword 1)) then returnm None
else
- (or_boolM
- (returnm ((neq_vec (privLevel_to_bits priv) ((privLevel_to_bits Machine) : mword 2))
- : bool))
- ((and_boolM
- (returnm ((eq_vec (privLevel_to_bits priv) ((privLevel_to_bits Machine) : mword 2))
- : bool))
+ (or_boolM (returnm ((generic_neq priv Machine) : bool))
+ ((and_boolM (returnm ((generic_eq priv Machine) : bool))
(read_reg mstatus_ref >>= fun w__3 : Mstatus =>
- returnm ((eq_vec (_get_Mstatus_MIE w__3) ((bool_to_bits true) : mword 1))
- : bool)))
+ returnm ((eq_vec (_get_Mstatus_MIE w__3) ('b"1" : mword 1)) : bool)))
: M (bool))) >>= fun mIE =>
(and_boolM ((haveSupMode tt) : M (bool))
- ((or_boolM
- (returnm ((eq_vec (privLevel_to_bits priv) ((privLevel_to_bits User) : mword 2))
- : bool))
- ((and_boolM
- (returnm ((eq_vec (privLevel_to_bits priv)
- ((privLevel_to_bits Supervisor)
- : mword 2))
- : bool))
+ ((or_boolM (returnm ((generic_eq priv User) : bool))
+ ((and_boolM (returnm ((generic_eq priv Supervisor) : bool))
(read_reg mstatus_ref >>= fun w__6 : Mstatus =>
- returnm ((eq_vec (_get_Mstatus_SIE w__6) ((bool_to_bits true) : mword 1))
- : bool)))
+ returnm ((eq_vec (_get_Mstatus_SIE w__6) ('b"1" : mword 1)) : bool)))
: M (bool)))
: M (bool))) >>= fun sIE =>
(and_boolM ((haveNExt tt) : M (bool))
- ((and_boolM
- (returnm ((eq_vec (privLevel_to_bits priv) ((privLevel_to_bits User) : mword 2))
- : bool))
+ ((and_boolM (returnm ((generic_eq priv User) : bool))
(read_reg mstatus_ref >>= fun w__10 : Mstatus =>
- returnm ((eq_vec (_get_Mstatus_UIE w__10) ((bool_to_bits true) : mword 1))
- : bool)))
+ returnm ((eq_vec (_get_Mstatus_UIE w__10) ('b"1" : mword 1)) : bool)))
: M (bool))) >>= fun uIE =>
read_reg mip_ref >>= fun w__12 : Minterrupts =>
read_reg mie_ref >>= fun w__13 : Minterrupts =>
read_reg mideleg_ref >>= fun w__14 : Minterrupts =>
(match (processPending w__12 w__13 (_get_Minterrupts_bits w__14) mIE) with
- | Ints_Empty (tt) => returnm (None : option ((mword 64 * Privilege)))
- | Ints_Pending (p) =>
+ | Ints_Empty tt => returnm None
+ | Ints_Pending p =>
let r := (p, Machine) in
- returnm ((Some
- (r))
- : option ((mword 64 * Privilege)))
- | Ints_Delegated (d) =>
+ returnm (Some r)
+ | Ints_Delegated d =>
(haveSupMode tt) >>= fun w__15 : bool =>
- (if sumbool_of_bool ((negb w__15)) then
- returnm ((if sumbool_of_bool (uIE) then
- let r := (d, User) in
- Some
- (r)
- else None)
- : option ((mword 64 * Privilege)))
+ (if sumbool_of_bool (negb w__15) then
+ returnm (if sumbool_of_bool uIE then
+ let r := (d, User) in
+ Some r
+ else None)
else
read_reg mie_ref >>= fun w__16 : Minterrupts =>
read_reg sideleg_ref >>= fun w__17 : Sinterrupts =>
- returnm ((match (processPending (Mk_Minterrupts d) w__16 (_get_Sinterrupts_bits w__17)
- sIE) with
- | Ints_Empty (tt) => None
- | Ints_Pending (p) =>
- let r := (p, Supervisor) in
- Some
- (r)
- | Ints_Delegated (d) =>
- if sumbool_of_bool (uIE) then
- let r := (d, User) in
- Some
- (r)
- else None
- end)
- : option ((mword 64 * Privilege))))
+ returnm (match (processPending (Mk_Minterrupts d) w__16 (_get_Sinterrupts_bits w__17)
+ sIE) with
+ | Ints_Empty tt => None
+ | Ints_Pending p =>
+ let r := (p, Supervisor) in
+ Some r
+ | Ints_Delegated d =>
+ if sumbool_of_bool uIE then
+ let r := (d, User) in
+ Some r
+ else None
+ end))
: M (option ((mword 64 * Privilege)))
end)
: M (option ((mword 64 * Privilege))))
: M (option ((mword 64 * Privilege))).
-Definition dispatchInterrupt (priv : Privilege)
-: M (option ((InterruptType * Privilege))) :=
-
+Definition dispatchInterrupt (priv : Privilege) : M (option ((InterruptType * Privilege))) :=
(or_boolM ((haveUsrMode tt) >>= fun w__0 : bool => returnm ((negb w__0) : bool))
((and_boolM ((haveSupMode tt) >>= fun w__1 : bool => returnm ((negb w__1) : bool))
((haveNExt tt) >>= fun w__2 : bool => returnm ((negb w__2) : bool)))
: M (bool))) >>= fun w__4 : bool =>
- (if sumbool_of_bool (w__4) then
- assert_exp (eq_vec (privLevel_to_bits priv) ((privLevel_to_bits Machine) : mword 2)) "invalid current privilege" >>
+ (if sumbool_of_bool w__4 then
+ assert_exp (generic_eq priv Machine) "invalid current privilege" >>
read_reg mip_ref >>= fun w__5 : Minterrupts =>
read_reg mie_ref >>= fun w__6 : Minterrupts =>
let enabled_pending := and_vec (_get_Minterrupts_bits w__5) (_get_Minterrupts_bits w__6) in
- returnm ((match (findPendingInterrupt enabled_pending) with
- | Some (i) =>
- let r := (i, Machine) in
- Some
- (r)
- | None => None
- end)
- : option ((InterruptType * Privilege)))
+ returnm (match (findPendingInterrupt enabled_pending) with
+ | Some i =>
+ let r := (i, Machine) in
+ Some r
+ | None => None
+ end)
else
(getPendingSet priv) >>= fun w__7 : option ((mword 64 * Privilege)) =>
- returnm ((match w__7 with
- | None => None
- | Some ((ip, p)) =>
- match (findPendingInterrupt ip) with
- | None => None
- | Some (i) =>
- let r := (i, p) in
- Some
- (r)
- end
- end)
- : option ((InterruptType * Privilege))))
+ returnm (match w__7 with
+ | None => None
+ | Some (ip, p) =>
+ match (findPendingInterrupt ip) with
+ | None => None
+ | Some i =>
+ let r := (i, p) in
+ Some r
+ end
+ end))
: M (option ((InterruptType * Privilege))).
-Definition tval (excinfo : option (mword 64))
-: mword 64 :=
-
- match excinfo with | Some (e) => e | None => EXTZ 64 (vec_of_bits [B0] : mword 1) end.
+Definition tval (excinfo : option (mword 64)) : mword 64 :=
+ match excinfo with | Some e => e | None => EXTZ 64 ('b"0" : mword 1) end.
-Definition rvfi_trap '(tt : unit) : unit := tt.
+Definition rvfi_trap '(tt : unit) : unit := tt.
Definition trap_handler
(del_priv : Privilege) (intr : bool) (c : mword 8) (pc : mword 64) (info : option (mword 64))
-(ext : option unit)
+(ext : option unit)
: M (mword 64) :=
-
let '_ := (rvfi_trap tt) : unit in
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "handling "
- (String.append (if sumbool_of_bool (intr) then "int#" else "exc#")
+ (String.append (if sumbool_of_bool intr then "int#" else "exc#")
(String.append (string_of_bits c)
(String.append " at priv "
(String.append (privLevel_to_str del_priv)
@@ -10993,38 +10904,37 @@ Definition trap_handler
let '_ := (cancel_reservation tt) : unit in
(match del_priv with
| Machine =>
- (_set_Mcause_IsInterrupt mcause_ref ((bool_to_bits intr) : mword 1)) >>
+ (_set_Mcause_IsInterrupt mcause_ref (bool_to_bits intr)) >>
(_set_Mcause_Cause mcause_ref (EXTZ 63 c)) >>
read_reg mstatus_ref >>= fun w__0 : Mstatus =>
(_set_Mstatus_MPIE mstatus_ref (_get_Mstatus_MIE w__0)) >>
- (_set_Mstatus_MIE mstatus_ref ((bool_to_bits false) : mword 1)) >>
+ (_set_Mstatus_MIE mstatus_ref ('b"0" : mword 1)) >>
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
(_set_Mstatus_MPP mstatus_ref (privLevel_to_bits w__1)) >>
write_reg mtval_ref (tval info) >>
write_reg mepc_ref pc >>
write_reg cur_privilege_ref del_priv >>
let '_ := (handle_trap_extension del_priv pc ext) : unit in
- (if ((get_config_print_reg tt)) then
+ (if get_config_print_reg tt then
read_reg mstatus_ref >>= fun w__2 : Mstatus =>
- returnm ((print_endline
- (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__2))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__2))))
+ else returnm tt) >>
read_reg mcause_ref >>= fun w__3 : Mcause =>
(prepare_trap_vector del_priv w__3)
: M (mword 64)
| Supervisor =>
(haveSupMode tt) >>= fun w__5 : bool =>
assert_exp' w__5 "no supervisor mode present for delegation" >>= fun _ =>
- (_set_Mcause_IsInterrupt scause_ref ((bool_to_bits intr) : mword 1)) >>
+ (_set_Mcause_IsInterrupt scause_ref (bool_to_bits intr)) >>
(_set_Mcause_Cause scause_ref (EXTZ 63 c)) >>
read_reg mstatus_ref >>= fun w__6 : Mstatus =>
(_set_Mstatus_SPIE mstatus_ref (_get_Mstatus_SIE w__6)) >>
- (_set_Mstatus_SIE mstatus_ref ((bool_to_bits false) : mword 1)) >>
+ (_set_Mstatus_SIE mstatus_ref ('b"0" : mword 1)) >>
read_reg cur_privilege_ref >>= fun w__7 : Privilege =>
(match w__7 with
- | User => returnm ((bool_to_bits false) : mword 1)
- | Supervisor => returnm ((bool_to_bits true) : mword 1)
+ | User => returnm ('b"0" : mword 1)
+ | Supervisor => returnm ('b"1" : mword 1)
| Machine => (internal_error "invalid privilege for s-mode trap") : M (mword 1)
end) >>= fun w__9 : mword 1 =>
(_set_Mstatus_SPP mstatus_ref w__9) >>
@@ -11032,47 +10942,44 @@ Definition trap_handler
write_reg sepc_ref pc >>
write_reg cur_privilege_ref del_priv >>
let '_ := (handle_trap_extension del_priv pc ext) : unit in
- (if ((get_config_print_reg tt)) then
+ (if get_config_print_reg tt then
read_reg mstatus_ref >>= fun w__10 : Mstatus =>
- returnm ((print_endline
- (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__10))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__10))))
+ else returnm tt) >>
read_reg scause_ref >>= fun w__11 : Mcause =>
(prepare_trap_vector del_priv w__11)
: M (mword 64)
| User =>
(haveUsrMode tt) >>= fun w__13 : bool =>
assert_exp' w__13 "no user mode present for delegation" >>= fun _ =>
- (_set_Mcause_IsInterrupt ucause_ref ((bool_to_bits intr) : mword 1)) >>
+ (_set_Mcause_IsInterrupt ucause_ref (bool_to_bits intr)) >>
(_set_Mcause_Cause ucause_ref (EXTZ 63 c)) >>
read_reg mstatus_ref >>= fun w__14 : Mstatus =>
(_set_Mstatus_UPIE mstatus_ref (_get_Mstatus_UIE w__14)) >>
- (_set_Mstatus_UIE mstatus_ref ((bool_to_bits false) : mword 1)) >>
+ (_set_Mstatus_UIE mstatus_ref ('b"0" : mword 1)) >>
write_reg utval_ref (tval info) >>
write_reg uepc_ref pc >>
write_reg cur_privilege_ref del_priv >>
let '_ := (handle_trap_extension del_priv pc ext) : unit in
- (if ((get_config_print_reg tt)) then
+ (if get_config_print_reg tt then
read_reg mstatus_ref >>= fun w__15 : Mstatus =>
- returnm ((print_endline
- (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__15))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__15))))
+ else returnm tt) >>
read_reg ucause_ref >>= fun w__16 : Mcause =>
(prepare_trap_vector del_priv w__16)
: M (mword 64)
end)
: M (mword 64).
-Definition exception_handler (cur_priv : Privilege) (ctl : ctl_result) (pc : mword 64)
+Definition exception_handler (cur_priv : Privilege) (ctl : ctl_result) (pc : mword 64)
: M (mword 64) :=
-
(match (cur_priv, ctl) with
- | (_, CTL_TRAP (e)) =>
+ | (_, CTL_TRAP e) =>
(exception_delegatee e.(sync_exception_trap) cur_priv) >>= fun del_priv =>
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "trapping from "
(String.append (privLevel_to_str cur_priv)
@@ -11082,163 +10989,175 @@ Definition exception_handler (cur_priv : Privilege) (ctl : ctl_result) (pc : mwo
(exceptionType_to_str e.(sync_exception_trap)))))))
else tt)
: unit in
- (trap_handler del_priv false ((exceptionType_to_bits e.(sync_exception_trap)) : mword 8) pc
+ (trap_handler del_priv false (exceptionType_to_bits e.(sync_exception_trap)) pc
e.(sync_exception_excinfo) e.(sync_exception_ext))
: M (mword 64)
- | (_, CTL_MRET (tt)) =>
+ | (_, CTL_MRET tt) =>
read_reg cur_privilege_ref >>= fun prev_priv =>
read_reg mstatus_ref >>= fun w__1 : Mstatus =>
(_set_Mstatus_MIE mstatus_ref (_get_Mstatus_MPIE w__1)) >>
- (_set_Mstatus_MPIE mstatus_ref ((bool_to_bits true) : mword 1)) >>
+ (_set_Mstatus_MPIE mstatus_ref ('b"1" : mword 1)) >>
read_reg mstatus_ref >>= fun w__2 : Mstatus =>
(privLevel_of_bits (_get_Mstatus_MPP w__2)) >>= fun w__3 : Privilege =>
write_reg cur_privilege_ref w__3 >>
(haveUsrMode tt) >>= fun w__4 : bool =>
(_set_Mstatus_MPP mstatus_ref
- (privLevel_to_bits (if sumbool_of_bool (w__4) then User else Machine))) >>
- (if ((get_config_print_reg tt)) then
- read_reg mstatus_ref >>= fun w__5 : Mstatus =>
- returnm ((print_endline
- (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__5))))
- : unit)
- else returnm (tt : unit)) >>
- (if ((get_config_print_platform tt)) then
- read_reg cur_privilege_ref >>= fun w__6 : Privilege =>
- returnm ((print_endline
- (String.append "ret-ing from "
- (String.append (privLevel_to_str prev_priv)
- (String.append " to " (privLevel_to_str w__6)))))
- : unit)
- else returnm (tt : unit)) >>
+ (privLevel_to_bits (if sumbool_of_bool w__4 then User else Machine))) >>
+ read_reg cur_privilege_ref >>= fun w__5 : Privilege =>
+ (if generic_neq w__5 Machine then
+ (_set_Mstatus_MPRV mstatus_ref ('b"0" : mword 1))
+ : M (unit)
+ else returnm tt) >>
+ (if get_config_print_reg tt then
+ read_reg mstatus_ref >>= fun w__6 : Mstatus =>
+ returnm (print_endline
+ (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__6))))
+ else returnm tt) >>
+ (if get_config_print_platform tt then
+ read_reg cur_privilege_ref >>= fun w__7 : Privilege =>
+ returnm (print_endline
+ (String.append "ret-ing from "
+ (String.append (privLevel_to_str prev_priv)
+ (String.append " to " (privLevel_to_str w__7)))))
+ else returnm tt) >>
let '_ := (cancel_reservation tt) : unit in
- (prepare_xret_target Machine) >>= fun w__7 : mword 64 =>
- (pc_alignment_mask tt) >>= fun w__8 : mword 64 => returnm ((and_vec w__7 w__8) : mword 64)
- | (_, CTL_SRET (tt)) =>
+ (prepare_xret_target Machine) >>= fun w__8 : mword 64 =>
+ (pc_alignment_mask tt) >>= fun w__9 : mword 64 => returnm (and_vec w__8 w__9)
+ | (_, CTL_SRET tt) =>
read_reg cur_privilege_ref >>= fun prev_priv =>
- read_reg mstatus_ref >>= fun w__9 : Mstatus =>
- (_set_Mstatus_SIE mstatus_ref (_get_Mstatus_SPIE w__9)) >>
- (_set_Mstatus_SPIE mstatus_ref ((bool_to_bits true) : mword 1)) >>
read_reg mstatus_ref >>= fun w__10 : Mstatus =>
+ (_set_Mstatus_SIE mstatus_ref (_get_Mstatus_SPIE w__10)) >>
+ (_set_Mstatus_SPIE mstatus_ref ('b"1" : mword 1)) >>
+ read_reg mstatus_ref >>= fun w__11 : Mstatus =>
write_reg
cur_privilege_ref
- (if ((eq_vec (_get_Mstatus_SPP w__10) ((bool_to_bits true) : mword 1))) then Supervisor
+ (if eq_vec (_get_Mstatus_SPP w__11) ('b"1" : mword 1) then Supervisor
else User) >>
- (_set_Mstatus_SPP mstatus_ref ((bool_to_bits false) : mword 1)) >>
- (if ((get_config_print_reg tt)) then
- read_reg mstatus_ref >>= fun w__11 : Mstatus =>
- returnm ((print_endline
- (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__11))))
- : unit)
- else returnm (tt : unit)) >>
- (if ((get_config_print_platform tt)) then
- read_reg cur_privilege_ref >>= fun w__12 : Privilege =>
- returnm ((print_endline
- (String.append "ret-ing from "
- (String.append (privLevel_to_str prev_priv)
- (String.append " to " (privLevel_to_str w__12)))))
- : unit)
- else returnm (tt : unit)) >>
+ (_set_Mstatus_SPP mstatus_ref ('b"0" : mword 1)) >>
+ read_reg cur_privilege_ref >>= fun w__12 : Privilege =>
+ (if generic_neq w__12 Machine then
+ (_set_Mstatus_MPRV mstatus_ref ('b"0" : mword 1))
+ : M (unit)
+ else returnm tt) >>
+ (if get_config_print_reg tt then
+ read_reg mstatus_ref >>= fun w__13 : Mstatus =>
+ returnm (print_endline
+ (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__13))))
+ else returnm tt) >>
+ (if get_config_print_platform tt then
+ read_reg cur_privilege_ref >>= fun w__14 : Privilege =>
+ returnm (print_endline
+ (String.append "ret-ing from "
+ (String.append (privLevel_to_str prev_priv)
+ (String.append " to " (privLevel_to_str w__14)))))
+ else returnm tt) >>
let '_ := (cancel_reservation tt) : unit in
- (prepare_xret_target Supervisor) >>= fun w__13 : mword 64 =>
- (pc_alignment_mask tt) >>= fun w__14 : mword 64 =>
- returnm ((and_vec w__13 w__14)
- : mword 64)
- | (_, CTL_URET (tt)) =>
+ (prepare_xret_target Supervisor) >>= fun w__15 : mword 64 =>
+ (pc_alignment_mask tt) >>= fun w__16 : mword 64 => returnm (and_vec w__15 w__16)
+ | (_, CTL_URET tt) =>
read_reg cur_privilege_ref >>= fun prev_priv =>
- read_reg mstatus_ref >>= fun w__15 : Mstatus =>
- (_set_Mstatus_UIE mstatus_ref (_get_Mstatus_UPIE w__15)) >>
- (_set_Mstatus_UPIE mstatus_ref ((bool_to_bits true) : mword 1)) >>
+ read_reg mstatus_ref >>= fun w__17 : Mstatus =>
+ (_set_Mstatus_UIE mstatus_ref (_get_Mstatus_UPIE w__17)) >>
+ (_set_Mstatus_UPIE mstatus_ref ('b"1" : mword 1)) >>
write_reg cur_privilege_ref User >>
- (if ((get_config_print_reg tt)) then
- read_reg mstatus_ref >>= fun w__16 : Mstatus =>
- returnm ((print_endline
- (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__16))))
- : unit)
- else returnm (tt : unit)) >>
- (if ((get_config_print_platform tt)) then
- read_reg cur_privilege_ref >>= fun w__17 : Privilege =>
- returnm ((print_endline
- (String.append "ret-ing from "
- (String.append (privLevel_to_str prev_priv)
- (String.append " to " (privLevel_to_str w__17)))))
- : unit)
- else returnm (tt : unit)) >>
+ (if get_config_print_reg tt then
+ read_reg mstatus_ref >>= fun w__18 : Mstatus =>
+ returnm (print_endline
+ (String.append "CSR mstatus <- " (string_of_bits (_get_Mstatus_bits w__18))))
+ else returnm tt) >>
+ (if get_config_print_platform tt then
+ read_reg cur_privilege_ref >>= fun w__19 : Privilege =>
+ returnm (print_endline
+ (String.append "ret-ing from "
+ (String.append (privLevel_to_str prev_priv)
+ (String.append " to " (privLevel_to_str w__19)))))
+ else returnm tt) >>
let '_ := (cancel_reservation tt) : unit in
- (prepare_xret_target User) >>= fun w__18 : mword 64 =>
- (pc_alignment_mask tt) >>= fun w__19 : mword 64 =>
- returnm ((and_vec w__18 w__19)
- : mword 64)
+ (prepare_xret_target User) >>= fun w__20 : mword 64 =>
+ (pc_alignment_mask tt) >>= fun w__21 : mword 64 => returnm (and_vec w__20 w__21)
end)
: M (mword 64).
-Definition handle_mem_exception (addr : mword 64) (e : ExceptionType)
-: M (unit) :=
-
+Definition handle_mem_exception (addr : mword 64) (e : ExceptionType) : M (unit) :=
let t : sync_exception :=
{| sync_exception_trap := e;
- sync_exception_excinfo := (Some (addr));
+ sync_exception_excinfo := (Some addr);
sync_exception_ext := None |} in
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
((read_reg PC_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
- (exception_handler w__0 (CTL_TRAP (t)) w__1) >>= fun w__2 : mword 64 =>
+ (exception_handler w__0 (CTL_TRAP t) w__1) >>= fun w__2 : mword 64 =>
(set_next_pc w__2)
: M (unit).
-Definition handle_interrupt (i : InterruptType) (del_priv : Privilege)
-: M (unit) :=
-
+Definition handle_exception (e : ExceptionType) : M (unit) :=
+ let t : sync_exception :=
+ {| sync_exception_trap := e;
+ sync_exception_excinfo := None;
+ sync_exception_ext := None |} in
+ read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
+ ((read_reg PC_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ (exception_handler w__0 (CTL_TRAP t) w__1) >>= fun w__2 : mword 64 =>
+ (set_next_pc w__2)
+ : M (unit).
+
+Definition handle_interrupt (i : InterruptType) (del_priv : Privilege) : M (unit) :=
((read_reg PC_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- (trap_handler del_priv true ((interruptType_to_bits i) : mword 8) w__0 None None) >>= fun w__1 : mword 64 =>
+ (trap_handler del_priv true (interruptType_to_bits i) w__0 None None) >>= fun w__1 : mword 64 =>
(set_next_pc w__1)
: M (unit).
-Definition init_sys '(tt : unit)
-: M (unit) :=
-
+Definition init_sys '(tt : unit) : M (unit) :=
write_reg cur_privilege_ref Machine >>
- write_reg mhartid_ref (EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
+ write_reg mhartid_ref (EXTZ 64 ('b"0" : mword 1)) >>
(_set_Misa_MXL misa_ref (arch_to_bits RV64)) >>
- (_set_Misa_A misa_ref ((bool_to_bits true) : mword 1)) >>
- (_set_Misa_C misa_ref ((bool_to_bits (sys_enable_rvc tt)) : mword 1)) >>
- (_set_Misa_I misa_ref ((bool_to_bits true) : mword 1)) >>
- (_set_Misa_M misa_ref ((bool_to_bits true) : mword 1)) >>
- (_set_Misa_U misa_ref ((bool_to_bits true) : mword 1)) >>
- (_set_Misa_S misa_ref ((bool_to_bits true) : mword 1)) >>
+ (_set_Misa_A misa_ref ('b"1" : mword 1)) >>
+ (_set_Misa_C misa_ref (bool_to_bits (sys_enable_rvc tt))) >>
+ (_set_Misa_I misa_ref ('b"1" : mword 1)) >>
+ (_set_Misa_M misa_ref ('b"1" : mword 1)) >>
+ (_set_Misa_U misa_ref ('b"1" : mword 1)) >>
+ (_set_Misa_S misa_ref ('b"1" : mword 1)) >>
+ (_set_Misa_F misa_ref (bool_to_bits (sys_enable_fdext tt))) >>
+ (_set_Misa_D misa_ref (bool_to_bits (sys_enable_fdext tt))) >>
read_reg mstatus_ref >>= fun w__0 : Mstatus =>
read_reg misa_ref >>= fun w__1 : Misa =>
write_reg mstatus_ref (set_mstatus_SXL w__0 (_get_Misa_MXL w__1)) >>
read_reg mstatus_ref >>= fun w__2 : Mstatus =>
read_reg misa_ref >>= fun w__3 : Misa =>
write_reg mstatus_ref (set_mstatus_UXL w__2 (_get_Misa_MXL w__3)) >>
- (_set_Mstatus_SD mstatus_ref ((bool_to_bits false) : mword 1)) >>
- (_set_Minterrupts_bits mip_ref (EXTZ 64 (vec_of_bits [B0] : mword 1))) >>
- (_set_Minterrupts_bits mie_ref (EXTZ 64 (vec_of_bits [B0] : mword 1))) >>
- (_set_Minterrupts_bits mideleg_ref (EXTZ 64 (vec_of_bits [B0] : mword 1))) >>
- (_set_Medeleg_bits medeleg_ref (EXTZ 64 (vec_of_bits [B0] : mword 1))) >>
- (_set_Mtvec_bits mtvec_ref (EXTZ 64 (vec_of_bits [B0] : mword 1))) >>
- (_set_Mcause_bits mcause_ref (EXTZ 64 (vec_of_bits [B0] : mword 1))) >>
- write_reg mepc_ref (EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
- write_reg mtval_ref (EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
- write_reg mscratch_ref (EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
- write_reg mcycle_ref (EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
- write_reg mtime_ref (EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
- (_set_Counteren_bits mcounteren_ref (EXTZ 32 (vec_of_bits [B0] : mword 1))) >>
- write_reg minstret_ref (EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
+ (_set_Mstatus_SD mstatus_ref ('b"0" : mword 1)) >>
+ (_set_Minterrupts_bits mip_ref (EXTZ 64 ('b"0" : mword 1))) >>
+ (_set_Minterrupts_bits mie_ref (EXTZ 64 ('b"0" : mword 1))) >>
+ (_set_Minterrupts_bits mideleg_ref (EXTZ 64 ('b"0" : mword 1))) >>
+ (_set_Medeleg_bits medeleg_ref (EXTZ 64 ('b"0" : mword 1))) >>
+ (_set_Mtvec_bits mtvec_ref (EXTZ 64 ('b"0" : mword 1))) >>
+ (_set_Mcause_bits mcause_ref (EXTZ 64 ('b"0" : mword 1))) >>
+ write_reg mepc_ref (EXTZ 64 ('b"0" : mword 1)) >>
+ write_reg mtval_ref (EXTZ 64 ('b"0" : mword 1)) >>
+ write_reg mscratch_ref (EXTZ 64 ('b"0" : mword 1)) >>
+ write_reg mcycle_ref (EXTZ 64 ('b"0" : mword 1)) >>
+ write_reg mtime_ref (EXTZ 64 ('b"0" : mword 1)) >>
+ (_set_Counteren_bits mcounteren_ref (EXTZ 32 ('b"0" : mword 1))) >>
+ write_reg minstret_ref (EXTZ 64 ('b"0" : mword 1)) >>
write_reg minstret_written_ref false >>
(init_pmp tt) >>
- (if ((get_config_print_reg tt)) then
+ (if get_config_print_reg tt then
read_reg mstatus_ref >>= fun w__4 : Mstatus =>
- returnm ((print_endline
- (String.append "CSR mstatus <- "
- (String.append (string_of_bits (_get_Mstatus_bits w__4))
- (String.append " (input: "
- (String.append
- (string_of_bits ((EXTZ 64 (vec_of_bits [B0] : mword 1)) : xlenbits))
- ")")))))
- : unit)
- else returnm (tt : unit))
+ returnm (print_endline
+ (String.append "CSR mstatus <- "
+ (String.append (string_of_bits (_get_Mstatus_bits w__4))
+ (String.append " (input: "
+ (String.append (string_of_bits ((EXTZ 64 ('b"0" : mword 1)) : xlenbits))
+ ")")))))
+ else returnm tt)
: M (unit).
+Definition MemoryOpResult_add_meta {t : Type} (r : MemoryOpResult t) (m : unit)
+: MemoryOpResult ((t * unit)) :=
+ match r with | MemValue v => MemValue (v, m) | MemException e => MemException e end.
+
+Definition MemoryOpResult_drop_meta {t : Type} (r : MemoryOpResult ((t * unit))) : MemoryOpResult t :=
+ match r with | MemValue (v, m) => MemValue v | MemException e => MemException e end.
+
Axiom elf_tohost : forall (_ : unit) , Z.
Axiom elf_entry : forall (_ : unit) , Z.
@@ -11263,28 +11182,24 @@ Axiom plat_clint_base : forall (_ : unit) , mword 64.
Axiom plat_clint_size : forall (_ : unit) , mword 64.
-Definition plat_htif_tohost '(tt : unit) : mword 64 := to_bits 64 (elf_tohost tt).
+Definition plat_htif_tohost '(tt : unit) : mword 64 := to_bits 64 (elf_tohost tt).
-Definition phys_mem_segments '(tt : unit)
-: list ((mword 64 * mword 64)) :=
-
+Definition phys_mem_segments '(tt : unit) : list ((mword 64 * mword 64)) :=
(plat_rom_base tt, plat_rom_size tt) :: (plat_ram_base tt, plat_ram_size tt) :: [].
-Definition within_phys_mem (addr : mword 64) (width : Z) `{ArithFact (width <= 16)}
-: bool :=
-
+Definition within_phys_mem (addr : mword 64) (width : Z) `{ArithFact (width <=? 16)} : bool :=
let addr_int := projT1 (uint addr) in
let ram_base_int := projT1 (uint (plat_ram_base tt)) in
let rom_base_int := projT1 (uint (plat_rom_base tt)) in
let ram_size_int := projT1 (uint (plat_ram_size tt)) in
let rom_size_int := projT1 (uint (plat_rom_size tt)) in
- if sumbool_of_bool ((andb (Z.leb ram_base_int addr_int)
- (Z.leb (Z.add addr_int (projT1 (__id width)))
- (Z.add ram_base_int ram_size_int)))) then
+ if sumbool_of_bool
+ (andb (Z.leb ram_base_int addr_int)
+ (Z.leb (Z.add addr_int (projT1 (__id width))) (Z.add ram_base_int ram_size_int))) then
true
- else if sumbool_of_bool ((andb (Z.leb rom_base_int addr_int)
- (Z.leb (Z.add addr_int (projT1 (__id width)))
- (Z.add rom_base_int rom_size_int)))) then
+ else if sumbool_of_bool
+ (andb (Z.leb rom_base_int addr_int)
+ (Z.leb (Z.add addr_int (projT1 (__id width))) (Z.add rom_base_int rom_size_int))) then
true
else
let '_ :=
@@ -11306,9 +11221,8 @@ Definition within_phys_mem (addr : mword 64) (width : Z) `{ArithFact (width <= 1
: unit in
false.
-Definition within_clint (addr : mword 64) (width : Z) `{ArithFact (0 < width /\ width <= 16)}
+Definition within_clint (addr : mword 64) (width : Z) `{ArithFact ((0 <? width) && (width <=? 16))}
: bool :=
-
let addr_int := projT1 (uint addr) in
let clint_base_int := projT1 (uint (plat_clint_base tt)) in
let clint_size_int := projT1 (uint (plat_clint_size tt)) in
@@ -11316,274 +11230,225 @@ Definition within_clint (addr : mword 64) (width : Z) `{ArithFact (0 < width /\
(Z.leb (Z.add addr_int (projT1 (__id width))) (Z.add clint_base_int clint_size_int)).
Definition within_htif_writable (addr : mword 64) (width : Z)
-`{ArithFact (0 < width /\ width <= 16)}
+`{ArithFact ((0 <? width) && (width <=? 16))}
: bool :=
-
orb (eq_vec (plat_htif_tohost tt) addr)
(andb (eq_vec (add_vec_int (plat_htif_tohost tt) 4) addr) (Z.eqb width 4)).
Definition within_htif_readable (addr : mword 64) (width : Z)
-`{ArithFact (0 < width /\ width <= 16)}
+`{ArithFact ((0 <? width) && (width <=? 16))}
: bool :=
-
orb (eq_vec (plat_htif_tohost tt) addr)
(andb (eq_vec (add_vec_int (plat_htif_tohost tt) 4) addr) (Z.eqb width 4)).
Axiom plat_insns_per_tick : forall (_ : unit) , Z.
-Definition MSIP_BASE : xlenbits :=
-EXTZ 64 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 20).
+Definition MSIP_BASE : xlenbits := EXTZ 64 (Ox"00000" : mword 20).
Hint Unfold MSIP_BASE : sail.
-Definition MTIMECMP_BASE : xlenbits :=
-EXTZ 64 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 20).
+Definition MTIMECMP_BASE : xlenbits := EXTZ 64 (Ox"04000" : mword 20).
Hint Unfold MTIMECMP_BASE : sail.
-Definition MTIMECMP_BASE_HI : xlenbits :=
-EXTZ 64 (vec_of_bits [B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0] : mword 20).
+Definition MTIMECMP_BASE_HI : xlenbits := EXTZ 64 (Ox"04004" : mword 20).
Hint Unfold MTIMECMP_BASE_HI : sail.
-Definition MTIME_BASE : xlenbits :=
-EXTZ 64 (vec_of_bits [B0;B0;B0;B0;B1;B0;B1;B1;B1;B1;B1;B1;B1;B1;B1;B1;B1;B0;B0;B0] : mword 20).
+Definition MTIME_BASE : xlenbits := EXTZ 64 (Ox"0BFF8" : mword 20).
Hint Unfold MTIME_BASE : sail.
-Definition MTIME_BASE_HI : xlenbits :=
-EXTZ 64 (vec_of_bits [B0;B0;B0;B0;B1;B0;B1;B1;B1;B1;B1;B1;B1;B1;B1;B1;B1;B1;B0;B0] : mword 20).
+Definition MTIME_BASE_HI : xlenbits := EXTZ 64 (Ox"0BFFC" : mword 20).
Hint Unfold MTIME_BASE_HI : sail.
-Definition clint_load (addr : mword 64) (width : Z) `{ArithFact (width > 0)}
+Definition clint_load (t : AccessType unit) (addr : mword 64) (width : Z) `{ArithFact (width >? 0)}
: M (MemoryOpResult (mword (8 * width))) :=
-
let addr := sub_vec addr (plat_clint_base tt) in
- (if sumbool_of_bool ((andb (eq_vec addr MSIP_BASE)
- (orb (Z.eqb (projT1 (__id width)) 8) (Z.eqb (projT1 (__id width)) 4))))
- then
- (if ((get_config_print_platform tt)) then
+ (if sumbool_of_bool
+ (andb (eq_vec addr MSIP_BASE)
+ (orb (Z.eqb (projT1 (__id width)) 8) (Z.eqb (projT1 (__id width)) 4))) then
+ (if get_config_print_platform tt then
read_reg mip_ref >>= fun w__0 : Minterrupts =>
- returnm ((print_endline
- (String.append "clint["
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits (_get_Minterrupts_MSI w__0))))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "clint["
+ (String.append (string_of_bits addr)
+ (String.append "] -> " (string_of_bits (_get_Minterrupts_MSI w__0))))))
+ else returnm tt) >>
read_reg mip_ref >>= fun w__1 : Minterrupts =>
- returnm ((MemValue
- (zero_extend (_get_Minterrupts_MSI w__1) (Z.mul 8 (projT1 (__id width)))))
- : MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (eq_vec addr MTIMECMP_BASE) (Z.eqb (projT1 (__id width)) 4)))
- then
- (if ((get_config_print_platform tt)) then
+ returnm (MemValue (zero_extend (_get_Minterrupts_MSI w__1) (Z.mul 8 (projT1 (__id width)))))
+ else if sumbool_of_bool (andb (eq_vec addr MTIMECMP_BASE) (Z.eqb (projT1 (__id width)) 4)) then
+ (if get_config_print_platform tt then
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
- returnm ((print_endline
- (String.append "clint<4>["
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits (subrange_vec_dec w__2 31 0))))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "clint<4>["
+ (String.append (string_of_bits addr)
+ (String.append "] -> " (string_of_bits (subrange_vec_dec w__2 31 0))))))
+ else returnm tt) >>
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__3 : mword 64 =>
- returnm ((MemValue
- (zero_extend (subrange_vec_dec w__3 31 0) _))
- : MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (eq_vec addr MTIMECMP_BASE) (Z.eqb (projT1 (__id width)) 8)))
- then
- (if ((get_config_print_platform tt)) then
+ returnm (MemValue (zero_extend (subrange_vec_dec w__3 31 0) _))
+ else if sumbool_of_bool (andb (eq_vec addr MTIMECMP_BASE) (Z.eqb (projT1 (__id width)) 8)) then
+ (if get_config_print_platform tt then
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__4 : mword 64 =>
- returnm ((print_endline
- (String.append "clint<8>["
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits w__4)))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "clint<8>["
+ (String.append (string_of_bits addr)
+ (String.append "] -> " (string_of_bits w__4)))))
+ else returnm tt) >>
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__5 : mword 64 =>
- returnm ((MemValue
- (zero_extend w__5 _))
- : MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (eq_vec addr MTIMECMP_BASE_HI) (Z.eqb (projT1 (__id width)) 4)))
- then
- (if ((get_config_print_platform tt)) then
+ returnm (MemValue (zero_extend w__5 _))
+ else if sumbool_of_bool
+ (andb (eq_vec addr MTIMECMP_BASE_HI) (Z.eqb (projT1 (__id width)) 4)) then
+ (if get_config_print_platform tt then
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__6 : mword 64 =>
- returnm ((print_endline
- (String.append "clint-hi<4>["
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits (subrange_vec_dec w__6 63 32))))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "clint-hi<4>["
+ (String.append (string_of_bits addr)
+ (String.append "] -> " (string_of_bits (subrange_vec_dec w__6 63 32))))))
+ else returnm tt) >>
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__7 : mword 64 =>
- returnm ((MemValue
- (zero_extend (subrange_vec_dec w__7 63 32) _))
- : MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (eq_vec addr MTIME_BASE) (Z.eqb (projT1 (__id width)) 4))) then
- (if ((get_config_print_platform tt)) then
+ returnm (MemValue (zero_extend (subrange_vec_dec w__7 63 32) _))
+ else if sumbool_of_bool (andb (eq_vec addr MTIME_BASE) (Z.eqb (projT1 (__id width)) 4)) then
+ (if get_config_print_platform tt then
((read_reg mtime_ref) : M (mword 64)) >>= fun w__8 : mword 64 =>
- returnm ((print_endline
- (String.append "clint["
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits w__8)))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "clint["
+ (String.append (string_of_bits addr)
+ (String.append "] -> " (string_of_bits w__8)))))
+ else returnm tt) >>
((read_reg mtime_ref) : M (mword 64)) >>= fun w__9 : mword 64 =>
- returnm ((MemValue
- (zero_extend (subrange_vec_dec w__9 31 0) _))
- : MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (eq_vec addr MTIME_BASE) (Z.eqb (projT1 (__id width)) 8))) then
- (if ((get_config_print_platform tt)) then
+ returnm (MemValue (zero_extend (subrange_vec_dec w__9 31 0) _))
+ else if sumbool_of_bool (andb (eq_vec addr MTIME_BASE) (Z.eqb (projT1 (__id width)) 8)) then
+ (if get_config_print_platform tt then
((read_reg mtime_ref) : M (mword 64)) >>= fun w__10 : mword 64 =>
- returnm ((print_endline
- (String.append "clint["
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits w__10)))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "clint["
+ (String.append (string_of_bits addr)
+ (String.append "] -> " (string_of_bits w__10)))))
+ else returnm tt) >>
((read_reg mtime_ref) : M (mword 64)) >>= fun w__11 : mword 64 =>
- returnm ((MemValue
- (zero_extend w__11 _))
- : MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (eq_vec addr MTIME_BASE_HI) (Z.eqb (projT1 (__id width)) 4)))
- then
- (if ((get_config_print_platform tt)) then
+ returnm (MemValue (zero_extend w__11 _))
+ else if sumbool_of_bool (andb (eq_vec addr MTIME_BASE_HI) (Z.eqb (projT1 (__id width)) 4)) then
+ (if get_config_print_platform tt then
((read_reg mtime_ref) : M (mword 64)) >>= fun w__12 : mword 64 =>
- returnm ((print_endline
- (String.append "clint["
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits w__12)))))
- : unit)
- else returnm (tt : unit)) >>
+ returnm (print_endline
+ (String.append "clint["
+ (String.append (string_of_bits addr)
+ (String.append "] -> " (string_of_bits w__12)))))
+ else returnm tt) >>
((read_reg mtime_ref) : M (mword 64)) >>= fun w__13 : mword 64 =>
- returnm ((MemValue
- (zero_extend (subrange_vec_dec w__13 63 32) _))
- : MemoryOpResult (mword (8 * width)))
+ returnm (MemValue (zero_extend (subrange_vec_dec w__13 63 32) _))
else
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "clint[" (String.append (string_of_bits addr) "] -> <not-mapped>"))
else tt)
: unit in
- returnm ((MemException
- (E_Load_Access_Fault))
- : MemoryOpResult (mword (8 * width))))
+ returnm (match t with
+ | Execute tt => MemException (E_Fetch_Access_Fault tt)
+ | Read Data => MemException (E_Load_Access_Fault tt)
+ | _ => MemException (E_SAMO_Access_Fault tt)
+ end))
: M (MemoryOpResult (mword (8 * width))).
-Definition clint_dispatch '(tt : unit)
-: M (unit) :=
-
- (if ((get_config_print_platform tt)) then
+Definition clint_dispatch '(tt : unit) : M (unit) :=
+ (if get_config_print_platform tt then
((read_reg mtime_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- returnm ((print_endline (String.append "clint::tick mtime <- " (string_of_bits w__0)))
- : unit)
- else returnm (tt : unit)) >>
- (_set_Minterrupts_MTI mip_ref ((bool_to_bits false) : mword 1)) >>
+ returnm (print_endline (String.append "clint::tick mtime <- " (string_of_bits w__0)))
+ else returnm tt) >>
+ (_set_Minterrupts_MTI mip_ref ('b"0" : mword 1)) >>
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
((read_reg mtime_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
- (if ((zopz0zIzJ_u w__1 w__2)) then
- (if ((get_config_print_platform tt)) then
+ (if zopz0zIzJ_u w__1 w__2 then
+ (if get_config_print_platform tt then
((read_reg mtime_ref) : M (mword 64)) >>= fun w__3 : mword 64 =>
- returnm ((print_endline
- (String.append " clint timer pending at mtime " (string_of_bits w__3)))
- : unit)
- else returnm (tt : unit)) >>
- (_set_Minterrupts_MTI mip_ref ((bool_to_bits true) : mword 1))
+ returnm (print_endline
+ (String.append " clint timer pending at mtime " (string_of_bits w__3)))
+ else returnm tt) >>
+ (_set_Minterrupts_MTI mip_ref ('b"1" : mword 1))
: M (unit)
- else returnm (tt : unit))
+ else returnm tt)
: M (unit).
Definition clint_store (addr : mword 64) (width : Z) (data : mword (8 * width))
-`{ArithFact (width > 0)}
+`{ArithFact (width >? 0)}
: M (MemoryOpResult bool) :=
-
let addr := sub_vec addr (plat_clint_base tt) in
- (if sumbool_of_bool ((andb (eq_vec addr MSIP_BASE)
- (orb (Z.eqb (projT1 (__id width)) 8) (Z.eqb (projT1 (__id width)) 4))))
- then
- (if ((get_config_print_platform tt)) then
- (cast_unit_vec (access_vec_dec data 0)) >>= fun w__0 : mword 1 =>
- returnm ((print_endline
- (String.append "clint["
- (String.append (string_of_bits addr)
- (String.append "] <- "
- (String.append (string_of_bits data)
- (String.append " (mip.MSI <- "
- (String.append (string_of_bits w__0) ")")))))))
- : unit)
- else returnm (tt : unit)) >>
- (cast_unit_vec (access_vec_dec data 0)) >>= fun w__1 : mword 1 =>
- (_set_Minterrupts_MSI mip_ref
- ((bool_to_bits (eq_vec w__1 (vec_of_bits [B1] : mword 1)))
- : mword 1)) >>
- (clint_dispatch tt) >> returnm ((MemValue (true)) : MemoryOpResult bool)
- else if sumbool_of_bool ((andb (eq_vec addr MTIMECMP_BASE) (Z.eqb (projT1 (__id width)) 8)))
- then
+ (if sumbool_of_bool
+ (andb (eq_vec addr MSIP_BASE)
+ (orb (Z.eqb (projT1 (__id width)) 8) (Z.eqb (projT1 (__id width)) 4))) then
+ (if get_config_print_platform tt then
+ (string_of_bit (access_vec_dec data 0)) >>= fun w__0 : string =>
+ returnm (print_endline
+ (String.append "clint["
+ (String.append (string_of_bits addr)
+ (String.append "] <- "
+ (String.append (string_of_bits data)
+ (String.append " (mip.MSI <- " (String.append w__0 ")")))))))
+ else returnm tt) >>
+ (_set_Minterrupts_MSI mip_ref (vec_of_bits [access_vec_dec data 0] : mword 1)) >>
+ (clint_dispatch tt) >> returnm (MemValue true)
+ else if sumbool_of_bool (andb (eq_vec addr MTIMECMP_BASE) (Z.eqb (projT1 (__id width)) 8)) then
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "clint<8>["
(String.append (string_of_bits addr)
(String.append "] <- " (String.append (string_of_bits data) " (mtimecmp)"))))
else tt)
: unit in
- write_reg mtimecmp_ref (zero_extend data 64) >>
- (clint_dispatch tt) >> returnm ((MemValue (true)) : MemoryOpResult bool)
- else if sumbool_of_bool ((andb (eq_vec addr MTIMECMP_BASE) (Z.eqb (projT1 (__id width)) 4)))
- then
+ write_reg mtimecmp_ref (zero_extend data 64) >> (clint_dispatch tt) >> returnm (MemValue true)
+ else if sumbool_of_bool (andb (eq_vec addr MTIMECMP_BASE) (Z.eqb (projT1 (__id width)) 4)) then
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "clint<4>["
(String.append (string_of_bits addr)
(String.append "] <- " (String.append (string_of_bits data) " (mtimecmp)"))))
else tt)
: unit in
- ((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
- write_reg mtimecmp_ref (update_subrange_vec_dec w__2 31 0 (zero_extend data 32)) >>
- (clint_dispatch tt) >> returnm ((MemValue (true)) : MemoryOpResult bool)
- else if sumbool_of_bool ((andb (eq_vec addr MTIMECMP_BASE_HI) (Z.eqb (projT1 (__id width)) 4)))
- then
+ ((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ write_reg mtimecmp_ref (update_subrange_vec_dec w__1 31 0 (zero_extend data 32)) >>
+ (clint_dispatch tt) >> returnm (MemValue true)
+ else if sumbool_of_bool
+ (andb (eq_vec addr MTIMECMP_BASE_HI) (Z.eqb (projT1 (__id width)) 4)) then
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "clint<4>["
(String.append (string_of_bits addr)
(String.append "] <- " (String.append (string_of_bits data) " (mtimecmp)"))))
else tt)
: unit in
- ((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__3 : mword 64 =>
- write_reg mtimecmp_ref (update_subrange_vec_dec w__3 63 32 (zero_extend data 32)) >>
- (clint_dispatch tt) >> returnm ((MemValue (true)) : MemoryOpResult bool)
+ ((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
+ write_reg mtimecmp_ref (update_subrange_vec_dec w__2 63 32 (zero_extend data 32)) >>
+ (clint_dispatch tt) >> returnm (MemValue true)
else
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "clint["
(String.append (string_of_bits addr)
(String.append "] <- " (String.append (string_of_bits data) " (<unmapped>)"))))
else tt)
: unit in
- returnm ((MemException
- (E_SAMO_Access_Fault))
- : MemoryOpResult bool))
+ returnm (MemException (E_SAMO_Access_Fault tt)))
: M (MemoryOpResult bool).
-Definition tick_clock '(tt : unit)
-: M (unit) :=
-
- ((read_reg mcycle_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- write_reg mcycle_ref (add_vec_int w__0 1) >>
- ((read_reg mtime_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
- write_reg mtime_ref (add_vec_int w__1 1) >> (clint_dispatch tt) : M (unit).
+Definition tick_clock '(tt : unit) : M (unit) :=
+ read_reg mcountinhibit_ref >>= fun w__0 : Counterin =>
+ (if eq_vec (_get_Counterin_CY w__0) ('b"0" : mword 1) then
+ ((read_reg mcycle_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
+ write_reg mcycle_ref (add_vec_int w__1 1)
+ : M (unit)
+ else returnm tt) >>
+ ((read_reg mtime_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
+ write_reg mtime_ref (add_vec_int w__2 1) >> (clint_dispatch tt) : M (unit).
Axiom plat_term_write : forall (_ : mword 8) , unit.
Axiom plat_term_read : forall (_ : unit) , mword 8.
-Definition Mk_htif_cmd (v : mword 64)
-: htif_cmd :=
-
+Definition Mk_htif_cmd (v : mword 64) : htif_cmd :=
{| htif_cmd_htif_cmd_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_htif_cmd_bits (v : htif_cmd)
-: mword 64 :=
-
+Definition _get_htif_cmd_bits (v : htif_cmd) : mword 64 :=
subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 63 0.
Definition _set_htif_cmd_bits (r_ref : register_ref regstate register_value htif_cmd) (v : mword 64)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -11593,22 +11458,17 @@ Definition _set_htif_cmd_bits (r_ref : register_ref regstate register_value htif
write_reg r_ref r
: M (unit).
-Definition _update_htif_cmd_bits (v : htif_cmd) (x : mword 64)
-: htif_cmd :=
-
+Definition _update_htif_cmd_bits (v : htif_cmd) (x : mword 64) : htif_cmd :=
{[ v with
htif_cmd_htif_cmd_chunk_0 :=
(update_subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_htif_cmd_device (v : htif_cmd)
-: mword 8 :=
-
+Definition _get_htif_cmd_device (v : htif_cmd) : mword 8 :=
subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 63 56.
Definition _set_htif_cmd_device
-(r_ref : register_ref regstate register_value htif_cmd) (v : mword 8)
+(r_ref : register_ref regstate register_value htif_cmd) (v : mword 8)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -11618,21 +11478,16 @@ Definition _set_htif_cmd_device
write_reg r_ref r
: M (unit).
-Definition _update_htif_cmd_device (v : htif_cmd) (x : mword 8)
-: htif_cmd :=
-
+Definition _update_htif_cmd_device (v : htif_cmd) (x : mword 8) : htif_cmd :=
{[ v with
htif_cmd_htif_cmd_chunk_0 :=
(update_subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 63 56 (subrange_vec_dec x 7 0)) ]}.
-Definition _get_htif_cmd_cmd (v : htif_cmd)
-: mword 8 :=
-
+Definition _get_htif_cmd_cmd (v : htif_cmd) : mword 8 :=
subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 55 48.
-Definition _set_htif_cmd_cmd (r_ref : register_ref regstate register_value htif_cmd) (v : mword 8)
+Definition _set_htif_cmd_cmd (r_ref : register_ref regstate register_value htif_cmd) (v : mword 8)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -11642,22 +11497,17 @@ Definition _set_htif_cmd_cmd (r_ref : register_ref regstate register_value htif_
write_reg r_ref r
: M (unit).
-Definition _update_htif_cmd_cmd (v : htif_cmd) (x : mword 8)
-: htif_cmd :=
-
+Definition _update_htif_cmd_cmd (v : htif_cmd) (x : mword 8) : htif_cmd :=
{[ v with
htif_cmd_htif_cmd_chunk_0 :=
(update_subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 55 48 (subrange_vec_dec x 7 0)) ]}.
-Definition _get_htif_cmd_payload (v : htif_cmd)
-: mword 48 :=
-
+Definition _get_htif_cmd_payload (v : htif_cmd) : mword 48 :=
subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 47 0.
Definition _set_htif_cmd_payload
-(r_ref : register_ref regstate register_value htif_cmd) (v : mword 48)
+(r_ref : register_ref regstate register_value htif_cmd) (v : mword 48)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -11667,61 +11517,55 @@ Definition _set_htif_cmd_payload
write_reg r_ref r
: M (unit).
-Definition _update_htif_cmd_payload (v : htif_cmd) (x : mword 48)
-: htif_cmd :=
-
+Definition _update_htif_cmd_payload (v : htif_cmd) (x : mword 48) : htif_cmd :=
{[ v with
htif_cmd_htif_cmd_chunk_0 :=
(update_subrange_vec_dec v.(htif_cmd_htif_cmd_chunk_0) 47 0 (subrange_vec_dec x 47 0)) ]}.
-Definition htif_load (addr : mword 64) (width : Z) `{ArithFact (width > 0)}
+Definition htif_load (t : AccessType unit) (paddr : mword 64) (width : Z) `{ArithFact (width >? 0)}
: M (MemoryOpResult (mword (8 * width))) :=
-
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
((read_reg htif_tohost_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- returnm ((print_endline
- (String.append "htif["
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits w__0)))))
- : unit)
- else returnm (tt : unit)) >>
- (if sumbool_of_bool ((andb (Z.eqb width 8) (eq_vec addr (plat_htif_tohost tt)))) then
+ returnm (print_endline
+ (String.append "htif["
+ (String.append (string_of_bits paddr)
+ (String.append "] -> " (string_of_bits w__0)))))
+ else returnm tt) >>
+ (if sumbool_of_bool (andb (Z.eqb width 8) (eq_vec paddr (plat_htif_tohost tt))) then
((read_reg htif_tohost_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
- returnm ((MemValue
- (zero_extend w__1 _))
- : MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (Z.eqb width 4) (eq_vec addr (plat_htif_tohost tt)))) then
+ returnm (MemValue (zero_extend w__1 _))
+ else if sumbool_of_bool (andb (Z.eqb width 4) (eq_vec paddr (plat_htif_tohost tt))) then
((read_reg htif_tohost_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
- returnm ((MemValue
- (zero_extend (subrange_vec_dec w__2 31 0) _))
- : MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (Z.eqb width 4)
- (eq_vec addr (add_vec_int (plat_htif_tohost tt) 4)))) then
+ returnm (MemValue (zero_extend (subrange_vec_dec w__2 31 0) _))
+ else if sumbool_of_bool
+ (andb (Z.eqb width 4) (eq_vec paddr (add_vec_int (plat_htif_tohost tt) 4))) then
((read_reg htif_tohost_ref) : M (mword 64)) >>= fun w__3 : mword 64 =>
- returnm ((MemValue
- (zero_extend (subrange_vec_dec w__3 63 32) _))
- : MemoryOpResult (mword (8 * width)))
- else returnm ((MemException (E_Load_Access_Fault)) : MemoryOpResult (mword (8 * width))))
+ returnm (MemValue (zero_extend (subrange_vec_dec w__3 63 32) _))
+ else
+ returnm (match t with
+ | Execute tt => MemException (E_Fetch_Access_Fault tt)
+ | Read Data => MemException (E_Load_Access_Fault tt)
+ | _ => MemException (E_SAMO_Access_Fault tt)
+ end))
: M (MemoryOpResult (mword (8 * width))).
-Definition htif_store (addr : mword 64) (width : Z) (data : mword (8 * width))
-`{ArithFact (0 < width /\ width <= 8)}
+Definition htif_store (paddr : mword 64) (width : Z) (data : mword (8 * width))
+`{ArithFact ((0 <? width) && (width <=? 8))}
: M (MemoryOpResult bool) :=
-
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "htif["
- (String.append (string_of_bits addr) (String.append "] <- " (string_of_bits data))))
+ (String.append (string_of_bits paddr) (String.append "] <- " (string_of_bits data))))
else tt)
: unit in
- (if sumbool_of_bool ((Z.eqb width 8)) then write_reg htif_tohost_ref (EXTZ 64 data) : M (unit)
- else if sumbool_of_bool ((andb (Z.eqb width 4) (eq_vec addr (plat_htif_tohost tt)))) then
+ (if sumbool_of_bool (Z.eqb width 8) then write_reg htif_tohost_ref (EXTZ 64 data) : M (unit)
+ else if sumbool_of_bool (andb (Z.eqb width 4) (eq_vec paddr (plat_htif_tohost tt))) then
((read_reg htif_tohost_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
write_reg htif_tohost_ref (update_subrange_vec_dec w__0 31 0 (autocast (autocast data)))
: M (unit)
- else if sumbool_of_bool ((andb (Z.eqb width 4)
- (eq_vec addr (add_vec_int (plat_htif_tohost tt) 4)))) then
+ else if sumbool_of_bool
+ (andb (Z.eqb width 4) (eq_vec paddr (add_vec_int (plat_htif_tohost tt) 4))) then
((read_reg htif_tohost_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
write_reg htif_tohost_ref (update_subrange_vec_dec w__1 63 32 (autocast (autocast data)))
: M (unit)
@@ -11729,417 +11573,358 @@ Definition htif_store (addr : mword 64) (width : Z) (data : mword (8 * width))
((read_reg htif_tohost_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
let cmd := Mk_htif_cmd w__2 in
let b__0 := _get_htif_cmd_device cmd in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8))) then
+ (if eq_vec b__0 (Ox"00" : mword 8) then
let '_ :=
- (if ((get_config_print_platform tt)) then
+ (if get_config_print_platform tt then
print_endline
(String.append "htif-syscall-proxy cmd: " (string_of_bits (_get_htif_cmd_payload cmd)))
else tt)
: unit in
- (cast_unit_vec (access_vec_dec (_get_htif_cmd_payload cmd) 0)) >>= fun w__3 : mword 1 =>
- (if ((eq_vec w__3 (vec_of_bits [B1] : mword 1))) then
+ (if eq_bit (access_vec_dec (_get_htif_cmd_payload cmd) 0) B1 then
write_reg htif_done_ref true >>
write_reg htif_exit_code_ref (shiftr (zero_extend (_get_htif_cmd_payload cmd) 64) 1)
: M (unit)
- else returnm (tt : unit))
+ else returnm tt)
: M (unit)
else
- returnm ((if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B1] : mword 8))) then
- let '_ :=
- (if ((get_config_print_platform tt)) then
- print_endline
- (String.append "htif-term cmd: "
- (string_of_bits (_get_htif_cmd_payload cmd)))
- else tt)
- : unit in
- let b__2 := _get_htif_cmd_cmd cmd in
- if ((eq_vec b__2 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8))) then tt
- else if ((eq_vec b__2 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B1] : mword 8))) then
- plat_term_write (subrange_vec_dec (_get_htif_cmd_payload cmd) 7 0)
- else print_endline (String.append "Unknown term cmd: " (string_of_bits b__2))
- else print_endline (String.append "htif-???? cmd: " (string_of_bits data)))
- : unit)) >>
- returnm ((MemValue
- (true))
- : MemoryOpResult bool).
-
-Definition htif_tick '(tt : unit)
-: M (unit) :=
-
- (if ((get_config_print_platform tt)) then
+ returnm (if eq_vec b__0 (Ox"01" : mword 8) then
+ let '_ :=
+ (if get_config_print_platform tt then
+ print_endline
+ (String.append "htif-term cmd: "
+ (string_of_bits (_get_htif_cmd_payload cmd)))
+ else tt)
+ : unit in
+ let b__2 := _get_htif_cmd_cmd cmd in
+ if eq_vec b__2 (Ox"00" : mword 8) then tt
+ else if eq_vec b__2 (Ox"01" : mword 8) then
+ plat_term_write (subrange_vec_dec (_get_htif_cmd_payload cmd) 7 0)
+ else print_endline (String.append "Unknown term cmd: " (string_of_bits b__2))
+ else print_endline (String.append "htif-???? cmd: " (string_of_bits data)))) >>
+ returnm (MemValue true).
+
+Definition htif_tick '(tt : unit) : M (unit) :=
+ (if get_config_print_platform tt then
((read_reg htif_tohost_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- returnm ((print_endline (String.append "htif::tick " (string_of_bits w__0)))
- : unit)
- else returnm (tt : unit)) >>
- write_reg htif_tohost_ref (EXTZ 64 (vec_of_bits [B0] : mword 1))
+ returnm (print_endline (String.append "htif::tick " (string_of_bits w__0)))
+ else returnm tt) >>
+ write_reg htif_tohost_ref (EXTZ 64 ('b"0" : mword 1))
: M (unit).
Definition within_mmio_readable (addr : mword 64) (width : Z)
-`{ArithFact (0 < width /\ width <= 16)}
+`{ArithFact ((0 <? width) && (width <=? 16))}
: bool :=
-
orb (within_clint addr width)
(andb (within_htif_readable addr width) (Z.leb 1 (projT1 (__id width)))).
Definition within_mmio_writable (addr : mword 64) (width : Z)
-`{ArithFact (0 < width /\ width <= 16)}
+`{ArithFact ((0 <? width) && (width <=? 16))}
: bool :=
-
orb (within_clint addr width)
(andb (within_htif_writable addr width) (Z.leb (projT1 (__id width)) 8)).
-Definition mmio_read (addr : mword 64) (width : Z) `{ArithFact (0 < width /\ width <= 16)}
+Definition mmio_read (t : AccessType unit) (paddr : mword 64) (width : Z)
+`{ArithFact ((0 <? width) && (width <=? 16))}
: M (MemoryOpResult (mword (8 * width))) :=
-
- (if ((within_clint addr width)) then
- (clint_load addr width)
+ (if within_clint paddr width then
+ (clint_load t paddr width)
: M (MemoryOpResult (mword (8 * width)))
- else if sumbool_of_bool ((andb (within_htif_readable addr width) (Z.leb 1 (projT1 (__id width)))))
- then
- (htif_load addr width)
+ else if sumbool_of_bool
+ (andb (within_htif_readable paddr width) (Z.leb 1 (projT1 (__id width)))) then
+ (htif_load t paddr width)
: M (MemoryOpResult (mword (8 * width)))
- else returnm ((MemException (E_Load_Access_Fault)) : MemoryOpResult (mword (8 * width))))
+ else
+ returnm (match t with
+ | Execute tt => MemException (E_Fetch_Access_Fault tt)
+ | Read Data => MemException (E_Load_Access_Fault tt)
+ | _ => MemException (E_SAMO_Access_Fault tt)
+ end))
: M (MemoryOpResult (mword (8 * width))).
-Definition mmio_write (addr : mword 64) (width : Z) (data : mword (8 * width))
-`{ArithFact (0 < width /\ width <= 16)}
+Definition mmio_write (paddr : mword 64) (width : Z) (data : mword (8 * width))
+`{ArithFact ((0 <? width) && (width <=? 16))}
: M (MemoryOpResult bool) :=
-
- (if ((within_clint addr width)) then (clint_store addr width data) : M (MemoryOpResult bool)
- else if sumbool_of_bool ((andb (within_htif_writable addr width) (Z.leb (projT1 (__id width)) 8)))
- then
- (htif_store addr width data)
+ (if within_clint paddr width then (clint_store paddr width data) : M (MemoryOpResult bool)
+ else if sumbool_of_bool
+ (andb (within_htif_writable paddr width) (Z.leb (projT1 (__id width)) 8)) then
+ (htif_store paddr width data)
: M (MemoryOpResult bool)
- else returnm ((MemException (E_SAMO_Access_Fault)) : MemoryOpResult bool))
+ else returnm (MemException (E_SAMO_Access_Fault tt)))
: M (MemoryOpResult bool).
-Definition init_platform '(tt : unit)
-: M (unit) :=
-
- write_reg htif_tohost_ref (EXTZ 64 (vec_of_bits [B0] : mword 1)) >>
+Definition init_platform '(tt : unit) : M (unit) :=
+ write_reg htif_tohost_ref (EXTZ 64 ('b"0" : mword 1)) >>
write_reg htif_done_ref false >>
- write_reg htif_exit_code_ref (EXTZ 64 (vec_of_bits [B0] : mword 1))
+ write_reg htif_exit_code_ref (EXTZ 64 ('b"0" : mword 1))
: M (unit).
-Definition tick_platform '(tt : unit) : M (unit) := (htif_tick tt) : M (unit).
+Definition tick_platform '(tt : unit) : M (unit) := (htif_tick tt) : M (unit).
-Definition handle_illegal '(tt : unit)
-: M (unit) :=
-
- (if ((plat_mtval_has_illegal_inst_bits tt)) then
- ((read_reg instbits_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- returnm ((Some
- (w__0))
- : option (mword 64))
- else returnm (None : option (mword 64))) >>= fun info =>
+Definition handle_illegal '(tt : unit) : M (unit) :=
+ (if plat_mtval_has_illegal_inst_bits tt then
+ ((read_reg instbits_ref) : M (mword 64)) >>= fun w__0 : mword 64 => returnm (Some w__0)
+ else returnm None) >>= fun info =>
let t : sync_exception :=
- {| sync_exception_trap := E_Illegal_Instr;
+ {| sync_exception_trap := (E_Illegal_Instr tt);
sync_exception_excinfo := info;
sync_exception_ext := None |} in
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
((read_reg PC_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
- (exception_handler w__1 (CTL_TRAP (t)) w__2) >>= fun w__3 : mword 64 =>
+ (exception_handler w__1 (CTL_TRAP t) w__2) >>= fun w__3 : mword 64 =>
(set_next_pc w__3)
: M (unit).
-Definition platform_wfi '(tt : unit)
-: M (unit) :=
-
+Definition platform_wfi '(tt : unit) : M (unit) :=
let '_ := (cancel_reservation tt) : unit in
((read_reg mtime_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
- (if ((zopz0zI_u w__0 w__1)) then
+ (if zopz0zI_u w__0 w__1 then
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
write_reg mtime_ref w__2 >>
((read_reg mtimecmp_ref) : M (mword 64)) >>= fun w__3 : mword 64 =>
write_reg mcycle_ref w__3
: M (unit)
- else returnm (tt : unit))
+ else returnm tt)
: M (unit).
-Definition is_aligned_addr (addr : mword 64) (width : Z)
-: bool :=
-
+Definition is_aligned_addr (addr : mword 64) (width : Z) : bool :=
Z.eqb (projT1 (emod_with_eq (projT1 (uint addr)) width)) 0.
+Definition read_kind_of_flags (aq : bool) (rl : bool) (res : bool) : option read_kind :=
+ match (aq, rl, res) with
+ | (false, false, false) => Some Read_plain
+ | (true, false, false) => Some Read_RISCV_acquire
+ | (true, true, false) => Some Read_RISCV_strong_acquire
+ | (false, false, true) => Some Read_RISCV_reserved
+ | (true, false, true) => Some Read_RISCV_reserved_acquire
+ | (true, true, true) => Some Read_RISCV_reserved_strong_acquire
+ | (false, true, false) => None
+ | (false, true, true) => None
+ end.
+
Definition phys_mem_read
-(t : AccessType) (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
-`{ArithFact (0 < width /\ width <= 16)}
-: M (MemoryOpResult (mword (8 * width))) :=
-
- (match (aq, rl, res) with
- | (false, false, false) =>
- (read_ram Read_plain addr width) >>= fun w__0 : mword (8 * width) =>
- returnm ((Some
- (w__0))
- : option (mword (8 * width)))
- | (true, false, false) =>
- (read_ram Read_RISCV_acquire addr width) >>= fun w__1 : mword (8 * width) =>
- returnm ((Some
- (w__1))
- : option (mword (8 * width)))
- | (true, true, false) =>
- (read_ram Read_RISCV_strong_acquire addr width) >>= fun w__2 : mword (8 * width) =>
- returnm ((Some
- (w__2))
- : option (mword (8 * width)))
- | (false, false, true) =>
- (read_ram Read_RISCV_reserved addr width) >>= fun w__3 : mword (8 * width) =>
- returnm ((Some
- (w__3))
- : option (mword (8 * width)))
- | (true, false, true) =>
- (read_ram Read_RISCV_reserved_acquire addr width) >>= fun w__4 : mword (8 * width) =>
- returnm ((Some
- (w__4))
- : option (mword (8 * width)))
- | (true, true, true) =>
- (read_ram Read_RISCV_reserved_strong_acquire addr width) >>= fun w__5 : mword (8 * width) =>
- returnm ((Some
- (w__5))
- : option (mword (8 * width)))
- | (false, true, false) => returnm (None : option (mword (8 * width)))
- | (false, true, true) => returnm (None : option (mword (8 * width)))
- end) >>= fun w__6 : option (mword (8 * width)) =>
- let result := w__6 : option (bits (8 * width)) in
- returnm ((match (t, result) with
- | (Execute, None) => MemException (E_Fetch_Access_Fault)
- | (Read, None) => MemException (E_Load_Access_Fault)
- | (_, None) => MemException (E_SAMO_Access_Fault)
- | (_, Some (v)) =>
- let '_ :=
- (if ((get_config_print_mem tt)) then
- print_endline
- (String.append "mem["
- (String.append (accessType_to_str t)
- (String.append ","
- (String.append (string_of_bits addr)
- (String.append "] -> " (string_of_bits v))))))
- else tt)
- : unit in
- MemValue
- (v)
- end)
- : MemoryOpResult (mword (8 * width))).
+(t : AccessType unit) (paddr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
+(meta : bool) `{ArithFact ((0 <? width) && (width <=? 16))}
+: M (MemoryOpResult ((mword (8 * width) * unit))) :=
+ (match (read_kind_of_flags aq rl res) with
+ | Some rk =>
+ (read_ram rk paddr width meta) >>= fun w__0 : (mword (8 * width) * unit) =>
+ returnm (Some w__0)
+ | None => returnm None
+ end) >>= fun w__1 : option ((mword (8 * width) * unit)) =>
+ let result := w__1 : option ((bits (8 * width) * mem_meta)) in
+ returnm (match (t, result) with
+ | (Execute tt, None) => MemException (E_Fetch_Access_Fault tt)
+ | (Read Data, None) => MemException (E_Load_Access_Fault tt)
+ | (_, None) => MemException (E_SAMO_Access_Fault tt)
+ | (_, Some (v, m)) =>
+ let '_ :=
+ (if get_config_print_mem tt then
+ print_endline
+ (String.append "mem["
+ (String.append (accessType_to_str t)
+ (String.append ","
+ (String.append (string_of_bits paddr)
+ (String.append "] -> " (string_of_bits v))))))
+ else tt)
+ : unit in
+ MemValue (v, m)
+ end).
Definition checked_mem_read
-(t : AccessType) (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
-`{ArithFact (0 < width /\ width <= 16)}
-: M (MemoryOpResult (mword (8 * width))) :=
-
- (if ((within_mmio_readable addr width)) then
- (mmio_read addr width)
- : M (MemoryOpResult (mword (8 * width)))
- else if ((within_phys_mem addr width)) then
- (phys_mem_read t addr width aq rl res)
- : M (MemoryOpResult (mword (8 * width)))
- else returnm ((MemException (E_Load_Access_Fault)) : MemoryOpResult (mword (8 * width))))
- : M (MemoryOpResult (mword (8 * width))).
+(t : AccessType unit) (paddr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
+(meta : bool) `{ArithFact ((0 <? width) && (width <=? 16))}
+: M (MemoryOpResult ((mword (8 * width) * unit))) :=
+ (if within_mmio_readable paddr width then
+ (mmio_read t paddr width) >>= fun w__0 : MemoryOpResult (mword (8 * width)) =>
+ returnm (MemoryOpResult_add_meta w__0 default_meta)
+ else if within_phys_mem paddr width then
+ (phys_mem_read t paddr width aq rl res meta)
+ : M (MemoryOpResult ((mword (8 * width) * unit)))
+ else
+ returnm (match t with
+ | Execute tt => MemException (E_Fetch_Access_Fault tt)
+ | Read Data => MemException (E_Load_Access_Fault tt)
+ | _ => MemException (E_SAMO_Access_Fault tt)
+ end))
+ : M (MemoryOpResult ((mword (8 * width) * unit))).
Definition pmp_mem_read
-(t : AccessType) (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
-`{ArithFact (0 < width /\ width <= 16)}
-: M (MemoryOpResult (mword (8 * width))) :=
-
- (if ((negb (plat_enable_pmp tt))) then
- (checked_mem_read t addr width aq rl res)
- : M (MemoryOpResult (mword (8 * width)))
+(t : AccessType unit) (paddr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
+(meta : bool) `{ArithFact ((0 <? width) && (width <=? 16))}
+: M (MemoryOpResult ((mword (8 * width) * unit))) :=
+ (if negb (plat_enable_pmp tt) then
+ (checked_mem_read t paddr width aq rl res meta)
+ : M (MemoryOpResult ((mword (8 * width) * unit)))
else
read_reg mstatus_ref >>= fun w__1 : Mstatus =>
read_reg cur_privilege_ref >>= fun w__2 : Privilege =>
- (effectivePrivilege w__1 w__2) >>= fun w__3 : Privilege =>
- (pmpCheck addr width t w__3) >>= fun w__4 : option ExceptionType =>
+ (effectivePrivilege t w__1 w__2) >>= fun w__3 : Privilege =>
+ (pmpCheck paddr width t w__3) >>= fun w__4 : option ExceptionType =>
(match w__4 with
- | None => (checked_mem_read t addr width aq rl res) : M (MemoryOpResult (mword (8 * width)))
- | Some (e) => returnm ((MemException (e)) : MemoryOpResult (mword (8 * width)))
+ | None =>
+ (checked_mem_read t paddr width aq rl res meta)
+ : M (MemoryOpResult ((mword (8 * width) * unit)))
+ | Some e => returnm (MemException e)
end)
- : M (MemoryOpResult (mword (8 * width))))
- : M (MemoryOpResult (mword (8 * width))).
+ : M (MemoryOpResult ((mword (8 * width) * unit))))
+ : M (MemoryOpResult ((mword (8 * width) * unit))).
Definition rvfi_read (addr : mword 64) (width : Z) (value : MemoryOpResult (mword (8 * width)))
-`{ArithFact (width > 0)}
+`{ArithFact (width >? 0)}
: unit :=
-
tt.
-Definition mem_read
-(typ : AccessType) (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
-`{ArithFact (0 < width /\ width <= 16)}
-: M (MemoryOpResult (mword (8 * width))) :=
-
- (if sumbool_of_bool ((andb (orb aq res) (negb (is_aligned_addr addr width)))) then
- returnm ((MemException
- (E_Load_Addr_Align))
- : MemoryOpResult (mword (8 * width)))
+Definition mem_read_meta
+(typ : AccessType unit) (paddr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
+(meta : bool) `{ArithFact ((0 <? width) && (width <=? 16))}
+: M (MemoryOpResult ((mword (8 * width) * unit))) :=
+ (if sumbool_of_bool (andb (orb aq res) (negb (is_aligned_addr paddr width))) then
+ returnm (MemException (E_Load_Addr_Align tt))
else
(match (aq, rl, res) with
- | (false, true, false) => throw (Error_not_implemented ("load.rl"))
- | (false, true, true) => throw (Error_not_implemented ("lr.rl"))
+ | (false, true, false) => throw (Error_not_implemented "load.rl")
+ | (false, true, true) => throw (Error_not_implemented "lr.rl")
| (_, _, _) =>
- (pmp_mem_read typ addr width aq rl res) : M (MemoryOpResult (mword (8 * width)))
+ (pmp_mem_read typ paddr width aq rl res meta)
+ : M (MemoryOpResult ((mword (8 * width) * unit)))
end)
- : M (MemoryOpResult (mword (8 * width)))) >>= fun result : MemoryOpResult (bits (8 * width)) =>
- let '_ := (rvfi_read addr width result) : unit in
- returnm (result
- : MemoryOpResult (mword (8 * width))).
+ : M (MemoryOpResult ((mword (8 * width) * unit)))) >>= fun result : MemoryOpResult ((bits (8 * width) * mem_meta)) =>
+ let '_ := (rvfi_read paddr width (MemoryOpResult_drop_meta result)) : unit in
+ returnm result.
+
+Definition mem_read
+(typ : AccessType unit) (paddr : mword 64) (width : Z) (aq : bool) (rl : bool) (res : bool)
+`{ArithFact ((0 <? width) && (width <=? 16))}
+: M (MemoryOpResult (mword (8 * width))) :=
+ (mem_read_meta typ paddr width aq rl res false) >>= fun w__0 : MemoryOpResult ((mword (8 * width) * unit)) =>
+ returnm (MemoryOpResult_drop_meta w__0).
Definition mem_write_ea (addr : mword 64) (width : Z) (aq : bool) (rl : bool) (con : bool)
-`{ArithFact (0 < width /\ width <= 16)}
+`{ArithFact ((0 <? width) && (width <=? 16))}
: M (MemoryOpResult unit) :=
-
- (if sumbool_of_bool ((andb (orb rl con) (negb (is_aligned_addr addr width)))) then
- returnm ((MemException
- (E_SAMO_Addr_Align))
- : MemoryOpResult unit)
+ (if sumbool_of_bool (andb (orb rl con) (negb (is_aligned_addr addr width))) then
+ returnm (MemException (E_SAMO_Addr_Align tt))
else
(match (aq, rl, con) with
- | (false, false, false) =>
- (write_ram_ea Write_plain addr width) >> returnm ((MemValue (tt)) : MemoryOpResult unit)
+ | (false, false, false) => (write_ram_ea Write_plain addr width) >> returnm (MemValue tt)
| (false, true, false) =>
- (write_ram_ea Write_RISCV_release addr width) >>
- returnm ((MemValue
- (tt))
- : MemoryOpResult unit)
+ (write_ram_ea Write_RISCV_release addr width) >> returnm (MemValue tt)
| (false, false, true) =>
- (write_ram_ea Write_RISCV_conditional addr width) >>
- returnm ((MemValue
- (tt))
- : MemoryOpResult unit)
+ (write_ram_ea Write_RISCV_conditional addr width) >> returnm (MemValue tt)
| (false, true, true) =>
- (write_ram_ea Write_RISCV_conditional_release addr width) >>
- returnm ((MemValue
- (tt))
- : MemoryOpResult unit)
- | (true, false, false) => throw (Error_not_implemented ("store.aq"))
+ (write_ram_ea Write_RISCV_conditional_release addr width) >> returnm (MemValue tt)
+ | (true, false, false) => throw (Error_not_implemented "store.aq")
| (true, true, false) =>
- (write_ram_ea Write_RISCV_strong_release addr width) >>
- returnm ((MemValue
- (tt))
- : MemoryOpResult unit)
- | (true, false, true) => throw (Error_not_implemented ("sc.aq"))
+ (write_ram_ea Write_RISCV_strong_release addr width) >> returnm (MemValue tt)
+ | (true, false, true) => throw (Error_not_implemented "sc.aq")
| (true, true, true) =>
- (write_ram_ea Write_RISCV_conditional_strong_release addr width) >>
- returnm ((MemValue
- (tt))
- : MemoryOpResult unit)
+ (write_ram_ea Write_RISCV_conditional_strong_release addr width) >> returnm (MemValue tt)
end)
: M (MemoryOpResult unit))
: M (MemoryOpResult unit).
Definition rvfi_write (addr : mword 64) (width : Z) (value : mword (8 * width))
-`{ArithFact (width > 0)}
+`{ArithFact (width >? 0)}
: unit :=
-
tt.
Definition phys_mem_write
-(wk : write_kind) (addr : mword 64) (width : Z) (data : mword (8 * width)) (meta : unit)
-`{ArithFact (0 < width /\ width <= 16)}
+(wk : write_kind) (paddr : mword 64) (width : Z) (data : mword (8 * width)) (meta : unit)
+`{ArithFact ((0 <? width) && (width <=? 16))}
: M (MemoryOpResult bool) :=
-
- let '_ := (rvfi_write addr width data) : unit in
- (write_ram wk addr width data meta) >>= fun w__0 : bool =>
- let result := MemValue (w__0) in
+ let '_ := (rvfi_write paddr width data) : unit in
+ (write_ram wk paddr width data meta) >>= fun w__0 : bool =>
+ let result := MemValue w__0 in
let '_ :=
- (if ((get_config_print_mem tt)) then
+ (if get_config_print_mem tt then
print_endline
(String.append "mem["
- (String.append (string_of_bits addr) (String.append "] <- " (string_of_bits data))))
+ (String.append (string_of_bits paddr) (String.append "] <- " (string_of_bits data))))
else tt)
: unit in
- returnm (result
- : MemoryOpResult bool).
+ returnm result.
Definition checked_mem_write
-(wk : write_kind) (addr : mword 64) (width : Z) (data : mword (8 * width)) (meta : unit)
-`{ArithFact (0 < width /\ width <= 16)}
+(wk : write_kind) (paddr : mword 64) (width : Z) (data : mword (8 * width)) (meta : unit)
+`{ArithFact ((0 <? width) && (width <=? 16))}
: M (MemoryOpResult bool) :=
-
- (if ((within_mmio_writable addr width)) then
- (mmio_write addr width data)
+ (if within_mmio_writable paddr width then
+ (mmio_write paddr width data)
: M (MemoryOpResult bool)
- else if ((within_phys_mem addr width)) then
- (phys_mem_write wk addr width data meta)
+ else if within_phys_mem paddr width then
+ (phys_mem_write wk paddr width data meta)
: M (MemoryOpResult bool)
- else returnm ((MemException (E_SAMO_Access_Fault)) : MemoryOpResult bool))
+ else returnm (MemException (E_SAMO_Access_Fault tt)))
: M (MemoryOpResult bool).
Definition pmp_mem_write
-(wk : write_kind) (addr : mword 64) (width : Z) (data : mword (8 * width)) (meta : unit)
-`{ArithFact (0 < width /\ width <= 16)}
+(wk : write_kind) (paddr : mword 64) (width : Z) (data : mword (8 * width)) (ext_acc : unit)
+(meta : unit) `{ArithFact ((0 <? width) && (width <=? 16))}
: M (MemoryOpResult bool) :=
-
- (if ((negb (plat_enable_pmp tt))) then
- (checked_mem_write wk addr width data meta)
+ (if negb (plat_enable_pmp tt) then
+ (checked_mem_write wk paddr width data meta)
: M (MemoryOpResult bool)
else
+ let typ : AccessType ext_access_type := Write ext_acc in
read_reg mstatus_ref >>= fun w__1 : Mstatus =>
read_reg cur_privilege_ref >>= fun w__2 : Privilege =>
- (effectivePrivilege w__1 w__2) >>= fun w__3 : Privilege =>
- (pmpCheck addr width Write w__3) >>= fun w__4 : option ExceptionType =>
+ (effectivePrivilege typ w__1 w__2) >>= fun w__3 : Privilege =>
+ (pmpCheck paddr width typ w__3) >>= fun w__4 : option ExceptionType =>
(match w__4 with
- | None => (checked_mem_write wk addr width data meta) : M (MemoryOpResult bool)
- | Some (e) => returnm ((MemException (e)) : MemoryOpResult bool)
+ | None => (checked_mem_write wk paddr width data meta) : M (MemoryOpResult bool)
+ | Some e => returnm (MemException e)
end)
: M (MemoryOpResult bool))
: M (MemoryOpResult bool).
Definition mem_write_value_meta
-(addr : mword 64) (width : Z) (value : mword (8 * width)) (meta : unit) (aq : bool) (rl : bool)
-(con : bool) `{ArithFact (0 < width /\ width <= 16)}
+(paddr : mword 64) (width : Z) (value : mword (8 * width)) (ext_acc : unit) (meta : unit)
+(aq : bool) (rl : bool) (con : bool) `{ArithFact ((0 <? width) && (width <=? 16))}
: M (MemoryOpResult bool) :=
-
- let '_ := (rvfi_write addr width value) : unit in
- (if sumbool_of_bool ((andb (orb rl con) (negb (is_aligned_addr addr width)))) then
- returnm ((MemException
- (E_SAMO_Addr_Align))
- : MemoryOpResult bool)
+ let '_ := (rvfi_write paddr width value) : unit in
+ (if sumbool_of_bool (andb (orb rl con) (negb (is_aligned_addr paddr width))) then
+ returnm (MemException (E_SAMO_Addr_Align tt))
else
(match (aq, rl, con) with
| (false, false, false) =>
- (pmp_mem_write Write_plain addr width value meta) : M (MemoryOpResult bool)
+ (pmp_mem_write Write_plain paddr width value ext_acc meta) : M (MemoryOpResult bool)
| (false, true, false) =>
- (pmp_mem_write Write_RISCV_release addr width value meta) : M (MemoryOpResult bool)
+ (pmp_mem_write Write_RISCV_release paddr width value ext_acc meta)
+ : M (MemoryOpResult bool)
| (false, false, true) =>
- (pmp_mem_write Write_RISCV_conditional addr width value meta) : M (MemoryOpResult bool)
+ (pmp_mem_write Write_RISCV_conditional paddr width value ext_acc meta)
+ : M (MemoryOpResult bool)
| (false, true, true) =>
- (pmp_mem_write Write_RISCV_conditional_release addr width value meta)
+ (pmp_mem_write Write_RISCV_conditional_release paddr width value ext_acc meta)
: M (MemoryOpResult bool)
| (true, true, false) =>
- (pmp_mem_write Write_RISCV_strong_release addr width value meta)
+ (pmp_mem_write Write_RISCV_strong_release paddr width value ext_acc meta)
: M (MemoryOpResult bool)
| (true, true, true) =>
- (pmp_mem_write Write_RISCV_conditional_strong_release addr width value meta)
+ (pmp_mem_write Write_RISCV_conditional_strong_release paddr width value ext_acc meta)
: M (MemoryOpResult bool)
- | (true, false, false) => throw (Error_not_implemented ("store.aq"))
- | (true, false, true) => throw (Error_not_implemented ("sc.aq"))
+ | (true, false, false) => throw (Error_not_implemented "store.aq")
+ | (true, false, true) => throw (Error_not_implemented "sc.aq")
end)
: M (MemoryOpResult bool))
: M (MemoryOpResult bool).
Definition mem_write_value
-(addr : mword 64) (width : Z) (value : mword (8 * width)) (aq : bool) (rl : bool) (con : bool)
-`{ArithFact (0 < width /\ width <= 16)}
+(paddr : mword 64) (width : Z) (value : mword (8 * width)) (aq : bool) (rl : bool) (con : bool)
+`{ArithFact ((0 <? width) && (width <=? 16))}
: M (MemoryOpResult bool) :=
-
- (mem_write_value_meta addr width value default_meta aq rl con)
+ (mem_write_value_meta paddr width value default_write_acc default_meta aq rl con)
: M (MemoryOpResult bool).
-Definition PAGESIZE_BITS := 12.
-Hint Unfold PAGESIZE_BITS : sail.
-Definition Mk_PTE_Bits (v : mword 8)
-: PTE_Bits :=
-
+Definition Mk_PTE_Bits (v : mword 8) : PTE_Bits :=
{| PTE_Bits_PTE_Bits_chunk_0 := (subrange_vec_dec v 7 0) |}.
-Definition _get_PTE_Bits_bits (v : PTE_Bits)
-: mword 8 :=
-
+Definition _get_PTE_Bits_bits (v : PTE_Bits) : mword 8 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 7 0.
-Definition _set_PTE_Bits_bits (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 8)
+Definition _set_PTE_Bits_bits (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 8)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12149,21 +11934,16 @@ Definition _set_PTE_Bits_bits (r_ref : register_ref regstate register_value PTE_
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_bits (v : PTE_Bits) (x : mword 8)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_bits (v : PTE_Bits) (x : mword 8) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 7 0 (subrange_vec_dec x 7 0)) ]}.
-Definition _get_PTE_Bits_D (v : PTE_Bits)
-: mword 1 :=
-
+Definition _get_PTE_Bits_D (v : PTE_Bits) : mword 1 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 7 7.
-Definition _set_PTE_Bits_D (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
+Definition _set_PTE_Bits_D (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12173,21 +11953,16 @@ Definition _set_PTE_Bits_D (r_ref : register_ref regstate register_value PTE_Bit
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_D (v : PTE_Bits) (x : mword 1)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_D (v : PTE_Bits) (x : mword 1) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 7 7 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_PTE_Bits_A (v : PTE_Bits)
-: mword 1 :=
-
+Definition _get_PTE_Bits_A (v : PTE_Bits) : mword 1 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 6 6.
-Definition _set_PTE_Bits_A (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
+Definition _set_PTE_Bits_A (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12197,21 +11972,16 @@ Definition _set_PTE_Bits_A (r_ref : register_ref regstate register_value PTE_Bit
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_A (v : PTE_Bits) (x : mword 1)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_A (v : PTE_Bits) (x : mword 1) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 6 6 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_PTE_Bits_G (v : PTE_Bits)
-: mword 1 :=
-
+Definition _get_PTE_Bits_G (v : PTE_Bits) : mword 1 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 5 5.
-Definition _set_PTE_Bits_G (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
+Definition _set_PTE_Bits_G (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12221,21 +11991,16 @@ Definition _set_PTE_Bits_G (r_ref : register_ref regstate register_value PTE_Bit
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_G (v : PTE_Bits) (x : mword 1)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_G (v : PTE_Bits) (x : mword 1) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 5 5 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_PTE_Bits_U (v : PTE_Bits)
-: mword 1 :=
-
+Definition _get_PTE_Bits_U (v : PTE_Bits) : mword 1 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 4 4.
-Definition _set_PTE_Bits_U (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
+Definition _set_PTE_Bits_U (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12245,21 +12010,16 @@ Definition _set_PTE_Bits_U (r_ref : register_ref regstate register_value PTE_Bit
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_U (v : PTE_Bits) (x : mword 1)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_U (v : PTE_Bits) (x : mword 1) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 4 4 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_PTE_Bits_X (v : PTE_Bits)
-: mword 1 :=
-
+Definition _get_PTE_Bits_X (v : PTE_Bits) : mword 1 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 3 3.
-Definition _set_PTE_Bits_X (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
+Definition _set_PTE_Bits_X (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12269,21 +12029,16 @@ Definition _set_PTE_Bits_X (r_ref : register_ref regstate register_value PTE_Bit
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_X (v : PTE_Bits) (x : mword 1)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_X (v : PTE_Bits) (x : mword 1) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 3 3 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_PTE_Bits_W (v : PTE_Bits)
-: mword 1 :=
-
+Definition _get_PTE_Bits_W (v : PTE_Bits) : mword 1 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 2 2.
-Definition _set_PTE_Bits_W (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
+Definition _set_PTE_Bits_W (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12293,21 +12048,16 @@ Definition _set_PTE_Bits_W (r_ref : register_ref regstate register_value PTE_Bit
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_W (v : PTE_Bits) (x : mword 1)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_W (v : PTE_Bits) (x : mword 1) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 2 2 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_PTE_Bits_R (v : PTE_Bits)
-: mword 1 :=
-
+Definition _get_PTE_Bits_R (v : PTE_Bits) : mword 1 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 1 1.
-Definition _set_PTE_Bits_R (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
+Definition _set_PTE_Bits_R (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12317,21 +12067,16 @@ Definition _set_PTE_Bits_R (r_ref : register_ref regstate register_value PTE_Bit
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_R (v : PTE_Bits) (x : mword 1)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_R (v : PTE_Bits) (x : mword 1) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 1 1 (subrange_vec_dec x 0 0)) ]}.
-Definition _get_PTE_Bits_V (v : PTE_Bits)
-: mword 1 :=
-
+Definition _get_PTE_Bits_V (v : PTE_Bits) : mword 1 :=
subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 0 0.
-Definition _set_PTE_Bits_V (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
+Definition _set_PTE_Bits_V (r_ref : register_ref regstate register_value PTE_Bits) (v : mword 1)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12341,142 +12086,116 @@ Definition _set_PTE_Bits_V (r_ref : register_ref regstate register_value PTE_Bit
write_reg r_ref r
: M (unit).
-Definition _update_PTE_Bits_V (v : PTE_Bits) (x : mword 1)
-: PTE_Bits :=
-
+Definition _update_PTE_Bits_V (v : PTE_Bits) (x : mword 1) : PTE_Bits :=
{[ v with
PTE_Bits_PTE_Bits_chunk_0 :=
(update_subrange_vec_dec v.(PTE_Bits_PTE_Bits_chunk_0) 0 0 (subrange_vec_dec x 0 0)) ]}.
-Definition isPTEPtr (p : mword 8)
-: bool :=
-
+Definition isPTEPtr (p : mword 8) (ext : mword 10) : bool :=
let a := Mk_PTE_Bits p in
- andb (eq_vec (_get_PTE_Bits_R a) ((bool_to_bits false) : mword 1))
- (andb (eq_vec (_get_PTE_Bits_W a) ((bool_to_bits false) : mword 1))
- (eq_vec (_get_PTE_Bits_X a) ((bool_to_bits false) : mword 1))).
+ andb (eq_vec (_get_PTE_Bits_R a) ('b"0" : mword 1))
+ (andb (eq_vec (_get_PTE_Bits_W a) ('b"0" : mword 1))
+ (eq_vec (_get_PTE_Bits_X a) ('b"0" : mword 1))).
-Definition isInvalidPTE (p : mword 8)
-: bool :=
-
+Definition isInvalidPTE (p : mword 8) (ext : mword 10) : bool :=
let a := Mk_PTE_Bits p in
- orb (eq_vec (_get_PTE_Bits_V a) ((bool_to_bits false) : mword 1))
- (andb (eq_vec (_get_PTE_Bits_W a) ((bool_to_bits true) : mword 1))
- (eq_vec (_get_PTE_Bits_R a) ((bool_to_bits false) : mword 1))).
+ orb (eq_vec (_get_PTE_Bits_V a) ('b"0" : mword 1))
+ (andb (eq_vec (_get_PTE_Bits_W a) ('b"1" : mword 1))
+ (eq_vec (_get_PTE_Bits_R a) ('b"0" : mword 1))).
+
+Definition to_pte_check (b : bool) : PTE_Check :=
+ if sumbool_of_bool b then PTE_Check_Success tt else PTE_Check_Failure tt.
Definition checkPTEPermission
-(ac : AccessType) (priv : Privilege) (mxr : bool) (do_sum : bool) (p : PTE_Bits)
-: M (bool) :=
-
+(ac : AccessType unit) (priv : Privilege) (mxr : bool) (do_sum : bool) (p : PTE_Bits)
+(ext : mword 10) (ext_ptw : unit)
+: M (PTE_Check) :=
(match (ac, priv) with
- | (Read, User) =>
- returnm (andb (eq_vec (_get_PTE_Bits_U p) ((bool_to_bits true) : mword 1))
- (orb (eq_vec (_get_PTE_Bits_R p) ((bool_to_bits true) : mword 1))
- (andb (eq_vec (_get_PTE_Bits_X p) ((bool_to_bits true) : mword 1)) mxr)))
- | (Write, User) =>
- returnm ((andb (eq_vec (_get_PTE_Bits_U p) ((bool_to_bits true) : mword 1))
- (eq_vec (_get_PTE_Bits_W p) ((bool_to_bits true) : mword 1)))
- : bool)
- | (ReadWrite, User) =>
- returnm (andb (eq_vec (_get_PTE_Bits_U p) ((bool_to_bits true) : mword 1))
- (andb (eq_vec (_get_PTE_Bits_W p) ((bool_to_bits true) : mword 1))
- (orb (eq_vec (_get_PTE_Bits_R p) ((bool_to_bits true) : mword 1))
- (andb (eq_vec (_get_PTE_Bits_X p) ((bool_to_bits true) : mword 1)) mxr))))
- | (Execute, User) =>
- returnm ((andb (eq_vec (_get_PTE_Bits_U p) ((bool_to_bits true) : mword 1))
- (eq_vec (_get_PTE_Bits_X p) ((bool_to_bits true) : mword 1)))
- : bool)
- | (Read, Supervisor) =>
- returnm (andb (orb (eq_vec (_get_PTE_Bits_U p) ((bool_to_bits false) : mword 1)) do_sum)
- (orb (eq_vec (_get_PTE_Bits_R p) ((bool_to_bits true) : mword 1))
- (andb (eq_vec (_get_PTE_Bits_X p) ((bool_to_bits true) : mword 1)) mxr)))
- | (Write, Supervisor) =>
- returnm (andb (orb (eq_vec (_get_PTE_Bits_U p) ((bool_to_bits false) : mword 1)) do_sum)
- (eq_vec (_get_PTE_Bits_W p) ((bool_to_bits true) : mword 1)))
- | (ReadWrite, Supervisor) =>
- returnm (andb (orb (eq_vec (_get_PTE_Bits_U p) ((bool_to_bits false) : mword 1)) do_sum)
- (andb (eq_vec (_get_PTE_Bits_W p) ((bool_to_bits true) : mword 1))
- (orb (eq_vec (_get_PTE_Bits_R p) ((bool_to_bits true) : mword 1))
- (andb (eq_vec (_get_PTE_Bits_X p) ((bool_to_bits true) : mword 1)) mxr))))
- | (Execute, Supervisor) =>
- returnm ((andb (eq_vec (_get_PTE_Bits_U p) ((bool_to_bits false) : mword 1))
- (eq_vec (_get_PTE_Bits_X p) ((bool_to_bits true) : mword 1)))
- : bool)
- | (_, Machine) => (internal_error "m-mode mem perm check") : M (bool)
+ | (Read Data, User) =>
+ returnm (to_pte_check
+ (andb (eq_vec (_get_PTE_Bits_U p) ('b"1" : mword 1))
+ (orb (eq_vec (_get_PTE_Bits_R p) ('b"1" : mword 1))
+ (andb (eq_vec (_get_PTE_Bits_X p) ('b"1" : mword 1)) mxr))))
+ | (Write Data, User) =>
+ returnm (to_pte_check
+ (andb (eq_vec (_get_PTE_Bits_U p) ('b"1" : mword 1))
+ (eq_vec (_get_PTE_Bits_W p) ('b"1" : mword 1))))
+ | (ReadWrite Data, User) =>
+ returnm (to_pte_check
+ (andb (eq_vec (_get_PTE_Bits_U p) ('b"1" : mword 1))
+ (andb (eq_vec (_get_PTE_Bits_W p) ('b"1" : mword 1))
+ (orb (eq_vec (_get_PTE_Bits_R p) ('b"1" : mword 1))
+ (andb (eq_vec (_get_PTE_Bits_X p) ('b"1" : mword 1)) mxr)))))
+ | (Execute tt, User) =>
+ returnm (to_pte_check
+ (andb (eq_vec (_get_PTE_Bits_U p) ('b"1" : mword 1))
+ (eq_vec (_get_PTE_Bits_X p) ('b"1" : mword 1))))
+ | (Read Data, Supervisor) =>
+ returnm (to_pte_check
+ (andb (orb (eq_vec (_get_PTE_Bits_U p) ('b"0" : mword 1)) do_sum)
+ (orb (eq_vec (_get_PTE_Bits_R p) ('b"1" : mword 1))
+ (andb (eq_vec (_get_PTE_Bits_X p) ('b"1" : mword 1)) mxr))))
+ | (Write Data, Supervisor) =>
+ returnm (to_pte_check
+ (andb (orb (eq_vec (_get_PTE_Bits_U p) ('b"0" : mword 1)) do_sum)
+ (eq_vec (_get_PTE_Bits_W p) ('b"1" : mword 1))))
+ | (ReadWrite Data, Supervisor) =>
+ returnm (to_pte_check
+ (andb (orb (eq_vec (_get_PTE_Bits_U p) ('b"0" : mword 1)) do_sum)
+ (andb (eq_vec (_get_PTE_Bits_W p) ('b"1" : mword 1))
+ (orb (eq_vec (_get_PTE_Bits_R p) ('b"1" : mword 1))
+ (andb (eq_vec (_get_PTE_Bits_X p) ('b"1" : mword 1)) mxr)))))
+ | (Execute tt, Supervisor) =>
+ returnm (to_pte_check
+ (andb (eq_vec (_get_PTE_Bits_U p) ('b"0" : mword 1))
+ (eq_vec (_get_PTE_Bits_X p) ('b"1" : mword 1))))
+ | (_, Machine) => (internal_error "m-mode mem perm check") : M (PTE_Check)
end)
- : M (bool).
+ : M (PTE_Check).
-Definition update_PTE_Bits (p : PTE_Bits) (a : AccessType)
-: option PTE_Bits :=
-
+Definition update_PTE_Bits (p : PTE_Bits) (a : AccessType unit) (ext : mword 10)
+: option ((PTE_Bits * mword 10)) :=
let update_d :=
- andb (orb (generic_eq a Write) (generic_eq a ReadWrite))
- (eq_vec (_get_PTE_Bits_D p) ((bool_to_bits false) : mword 1)) in
- let update_a := eq_vec (_get_PTE_Bits_A p) ((bool_to_bits false) : mword 1) in
- if sumbool_of_bool ((orb update_d update_a)) then
- let np := _update_PTE_Bits_A p ((bool_to_bits true) : mword 1) in
- let np :=
- if sumbool_of_bool (update_d) then _update_PTE_Bits_D np ((bool_to_bits true) : mword 1)
- else np in
- Some
- (np)
- else None.
-
-Definition PTW_Error_of_num (arg_ : Z) `{ArithFact (0 <= arg_ /\ arg_ <= 4)}
-: PTW_Error :=
-
- let l__8 := arg_ in
- if sumbool_of_bool ((Z.eqb l__8 0)) then PTW_Access
- else if sumbool_of_bool ((Z.eqb l__8 1)) then PTW_Invalid_PTE
- else if sumbool_of_bool ((Z.eqb l__8 2)) then PTW_No_Permission
- else if sumbool_of_bool ((Z.eqb l__8 3)) then PTW_Misaligned
- else PTW_PTE_Update.
-
-Definition num_of_PTW_Error (arg_ : PTW_Error)
-: {e : Z & ArithFact (0 <= e /\ e <= 4)} :=
-
- build_ex(match arg_ with
- | PTW_Access => 0
- | PTW_Invalid_PTE => 1
- | PTW_No_Permission => 2
- | PTW_Misaligned => 3
- | PTW_PTE_Update => 4
- end).
+ andb (orb (generic_eq a (Write Data)) (generic_eq a (ReadWrite Data)))
+ (eq_vec (_get_PTE_Bits_D p) ('b"0" : mword 1)) in
+ let update_a := eq_vec (_get_PTE_Bits_A p) ('b"0" : mword 1) in
+ if sumbool_of_bool (orb update_d update_a) then
+ let np := _update_PTE_Bits_A p ('b"1" : mword 1) in
+ let np := if sumbool_of_bool update_d then _update_PTE_Bits_D np ('b"1" : mword 1) else np in
+ Some (np, ext)
+ else None.
-Definition ptw_error_to_str (e : PTW_Error)
-: string :=
-
+Definition ptw_error_to_str (e : PTW_Error) : string :=
match e with
- | PTW_Access => "mem-access-error"
- | PTW_Invalid_PTE => "invalid-pte"
- | PTW_No_Permission => "no-permission"
- | PTW_Misaligned => "misaligned-superpage"
- | PTW_PTE_Update => "pte-update-needed"
+ | PTW_Invalid_Addr tt => "invalid-source-addr"
+ | PTW_Access tt => "mem-access-error"
+ | PTW_Invalid_PTE tt => "invalid-pte"
+ | PTW_No_Permission tt => "no-permission"
+ | PTW_Misaligned tt => "misaligned-superpage"
+ | PTW_PTE_Update tt => "pte-update-needed"
+ | PTW_Ext_Error e => "extension-error"
end.
-Definition translationException (a : AccessType) (f : PTW_Error)
-: ExceptionType :=
-
+Definition translationException (a : AccessType unit) (f : PTW_Error) : ExceptionType :=
match (a, f) with
- | (ReadWrite, PTW_Access) => E_SAMO_Access_Fault
- | (ReadWrite, _) => E_SAMO_Page_Fault
- | (Read, PTW_Access) => E_Load_Access_Fault
- | (Read, _) => E_Load_Page_Fault
- | (Write, PTW_Access) => E_SAMO_Access_Fault
- | (Write, _) => E_SAMO_Page_Fault
- | (Fetch, PTW_Access) => E_Fetch_Access_Fault
- | (Fetch, _) => E_Fetch_Page_Fault
+ | (_, PTW_Ext_Error e) => E_Extension (ext_translate_exception e)
+ | (ReadWrite Data, PTW_Access tt) => E_SAMO_Access_Fault tt
+ | (ReadWrite Data, _) => E_SAMO_Page_Fault tt
+ | (Read Data, PTW_Access tt) => E_Load_Access_Fault tt
+ | (Read Data, _) => E_Load_Page_Fault tt
+ | (Write Data, PTW_Access tt) => E_SAMO_Access_Fault tt
+ | (Write Data, _) => E_SAMO_Page_Fault tt
+ | (Execute tt, PTW_Access tt) => E_Fetch_Access_Fault tt
+ | (Execute tt, _) => E_Fetch_Page_Fault tt
end.
-Definition curAsid32 (satp : mword 32)
-: mword 9 :=
-
+Definition PAGESIZE_BITS := 12.
+Hint Unfold PAGESIZE_BITS : sail.
+Definition curAsid32 (satp : mword 32) : mword 9 :=
let s := Mk_Satp32 satp in
_get_Satp32_Asid s.
-Definition curPTB32 (satp : mword 32)
-: mword 34 :=
-
+Definition curPTB32 (satp : mword 32) : mword 34 :=
let s : Satp32 := Mk_Satp32 satp in
shiftl (EXTZ 34 (_get_Satp32_PPN s)) PAGESIZE_BITS.
@@ -12488,20 +12207,15 @@ Definition PTE32_LOG_SIZE := 2.
Hint Unfold PTE32_LOG_SIZE : sail.
Definition PTE32_SIZE := 4.
Hint Unfold PTE32_SIZE : sail.
-Definition Mk_SV32_Vaddr (v : mword 32)
-: SV32_Vaddr :=
-
+Definition Mk_SV32_Vaddr (v : mword 32) : SV32_Vaddr :=
{| SV32_Vaddr_SV32_Vaddr_chunk_0 := (subrange_vec_dec v 31 0) |}.
-Definition _get_SV32_Vaddr_bits (v : SV32_Vaddr)
-: mword 32 :=
-
+Definition _get_SV32_Vaddr_bits (v : SV32_Vaddr) : mword 32 :=
subrange_vec_dec v.(SV32_Vaddr_SV32_Vaddr_chunk_0) 31 0.
Definition _set_SV32_Vaddr_bits
-(r_ref : register_ref regstate register_value SV32_Vaddr) (v : mword 32)
+(r_ref : register_ref regstate register_value SV32_Vaddr) (v : mword 32)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12511,22 +12225,17 @@ Definition _set_SV32_Vaddr_bits
write_reg r_ref r
: M (unit).
-Definition _update_SV32_Vaddr_bits (v : SV32_Vaddr) (x : mword 32)
-: SV32_Vaddr :=
-
+Definition _update_SV32_Vaddr_bits (v : SV32_Vaddr) (x : mword 32) : SV32_Vaddr :=
{[ v with
SV32_Vaddr_SV32_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV32_Vaddr_SV32_Vaddr_chunk_0) 31 0 (subrange_vec_dec x 31 0)) ]}.
-Definition _get_SV32_Vaddr_VPNi (v : SV32_Vaddr)
-: mword 20 :=
-
+Definition _get_SV32_Vaddr_VPNi (v : SV32_Vaddr) : mword 20 :=
subrange_vec_dec v.(SV32_Vaddr_SV32_Vaddr_chunk_0) 31 12.
Definition _set_SV32_Vaddr_VPNi
-(r_ref : register_ref regstate register_value SV32_Vaddr) (v : mword 20)
+(r_ref : register_ref regstate register_value SV32_Vaddr) (v : mword 20)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12536,22 +12245,17 @@ Definition _set_SV32_Vaddr_VPNi
write_reg r_ref r
: M (unit).
-Definition _update_SV32_Vaddr_VPNi (v : SV32_Vaddr) (x : mword 20)
-: SV32_Vaddr :=
-
+Definition _update_SV32_Vaddr_VPNi (v : SV32_Vaddr) (x : mword 20) : SV32_Vaddr :=
{[ v with
SV32_Vaddr_SV32_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV32_Vaddr_SV32_Vaddr_chunk_0) 31 12 (subrange_vec_dec x 19 0)) ]}.
-Definition _get_SV32_Vaddr_PgOfs (v : SV32_Vaddr)
-: mword 12 :=
-
+Definition _get_SV32_Vaddr_PgOfs (v : SV32_Vaddr) : mword 12 :=
subrange_vec_dec v.(SV32_Vaddr_SV32_Vaddr_chunk_0) 11 0.
Definition _set_SV32_Vaddr_PgOfs
-(r_ref : register_ref regstate register_value SV32_Vaddr) (v : mword 12)
+(r_ref : register_ref regstate register_value SV32_Vaddr) (v : mword 12)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12561,27 +12265,20 @@ Definition _set_SV32_Vaddr_PgOfs
write_reg r_ref r
: M (unit).
-Definition _update_SV32_Vaddr_PgOfs (v : SV32_Vaddr) (x : mword 12)
-: SV32_Vaddr :=
-
+Definition _update_SV32_Vaddr_PgOfs (v : SV32_Vaddr) (x : mword 12) : SV32_Vaddr :=
{[ v with
SV32_Vaddr_SV32_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV32_Vaddr_SV32_Vaddr_chunk_0) 11 0 (subrange_vec_dec x 11 0)) ]}.
-Definition Mk_SV32_Paddr (v : mword 34)
-: SV32_Paddr :=
-
+Definition Mk_SV32_Paddr (v : mword 34) : SV32_Paddr :=
{| SV32_Paddr_SV32_Paddr_chunk_0 := (subrange_vec_dec v 33 0) |}.
-Definition _get_SV32_Paddr_bits (v : SV32_Paddr)
-: mword 34 :=
-
+Definition _get_SV32_Paddr_bits (v : SV32_Paddr) : mword 34 :=
subrange_vec_dec v.(SV32_Paddr_SV32_Paddr_chunk_0) 33 0.
Definition _set_SV32_Paddr_bits
-(r_ref : register_ref regstate register_value SV32_Paddr) (v : mword 34)
+(r_ref : register_ref regstate register_value SV32_Paddr) (v : mword 34)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12591,22 +12288,17 @@ Definition _set_SV32_Paddr_bits
write_reg r_ref r
: M (unit).
-Definition _update_SV32_Paddr_bits (v : SV32_Paddr) (x : mword 34)
-: SV32_Paddr :=
-
+Definition _update_SV32_Paddr_bits (v : SV32_Paddr) (x : mword 34) : SV32_Paddr :=
{[ v with
SV32_Paddr_SV32_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV32_Paddr_SV32_Paddr_chunk_0) 33 0 (subrange_vec_dec x 33 0)) ]}.
-Definition _get_SV32_Paddr_PPNi (v : SV32_Paddr)
-: mword 22 :=
-
+Definition _get_SV32_Paddr_PPNi (v : SV32_Paddr) : mword 22 :=
subrange_vec_dec v.(SV32_Paddr_SV32_Paddr_chunk_0) 33 12.
Definition _set_SV32_Paddr_PPNi
-(r_ref : register_ref regstate register_value SV32_Paddr) (v : mword 22)
+(r_ref : register_ref regstate register_value SV32_Paddr) (v : mword 22)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12616,22 +12308,17 @@ Definition _set_SV32_Paddr_PPNi
write_reg r_ref r
: M (unit).
-Definition _update_SV32_Paddr_PPNi (v : SV32_Paddr) (x : mword 22)
-: SV32_Paddr :=
-
+Definition _update_SV32_Paddr_PPNi (v : SV32_Paddr) (x : mword 22) : SV32_Paddr :=
{[ v with
SV32_Paddr_SV32_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV32_Paddr_SV32_Paddr_chunk_0) 33 12 (subrange_vec_dec x 21 0)) ]}.
-Definition _get_SV32_Paddr_PgOfs (v : SV32_Paddr)
-: mword 12 :=
-
+Definition _get_SV32_Paddr_PgOfs (v : SV32_Paddr) : mword 12 :=
subrange_vec_dec v.(SV32_Paddr_SV32_Paddr_chunk_0) 11 0.
Definition _set_SV32_Paddr_PgOfs
-(r_ref : register_ref regstate register_value SV32_Paddr) (v : mword 12)
+(r_ref : register_ref regstate register_value SV32_Paddr) (v : mword 12)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12641,27 +12328,19 @@ Definition _set_SV32_Paddr_PgOfs
write_reg r_ref r
: M (unit).
-Definition _update_SV32_Paddr_PgOfs (v : SV32_Paddr) (x : mword 12)
-: SV32_Paddr :=
-
+Definition _update_SV32_Paddr_PgOfs (v : SV32_Paddr) (x : mword 12) : SV32_Paddr :=
{[ v with
SV32_Paddr_SV32_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV32_Paddr_SV32_Paddr_chunk_0) 11 0 (subrange_vec_dec x 11 0)) ]}.
-Definition Mk_SV32_PTE (v : mword 32)
-: SV32_PTE :=
-
+Definition Mk_SV32_PTE (v : mword 32) : SV32_PTE :=
{| SV32_PTE_SV32_PTE_chunk_0 := (subrange_vec_dec v 31 0) |}.
-Definition _get_SV32_PTE_bits (v : SV32_PTE)
-: mword 32 :=
-
+Definition _get_SV32_PTE_bits (v : SV32_PTE) : mword 32 :=
subrange_vec_dec v.(SV32_PTE_SV32_PTE_chunk_0) 31 0.
Definition _set_SV32_PTE_bits (r_ref : register_ref regstate register_value SV32_PTE) (v : mword 32)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12671,22 +12350,16 @@ Definition _set_SV32_PTE_bits (r_ref : register_ref regstate register_value SV32
write_reg r_ref r
: M (unit).
-Definition _update_SV32_PTE_bits (v : SV32_PTE) (x : mword 32)
-: SV32_PTE :=
-
+Definition _update_SV32_PTE_bits (v : SV32_PTE) (x : mword 32) : SV32_PTE :=
{[ v with
SV32_PTE_SV32_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV32_PTE_SV32_PTE_chunk_0) 31 0 (subrange_vec_dec x 31 0)) ]}.
-Definition _get_SV32_PTE_PPNi (v : SV32_PTE)
-: mword 22 :=
-
+Definition _get_SV32_PTE_PPNi (v : SV32_PTE) : mword 22 :=
subrange_vec_dec v.(SV32_PTE_SV32_PTE_chunk_0) 31 10.
Definition _set_SV32_PTE_PPNi (r_ref : register_ref regstate register_value SV32_PTE) (v : mword 22)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12696,21 +12369,16 @@ Definition _set_SV32_PTE_PPNi (r_ref : register_ref regstate register_value SV32
write_reg r_ref r
: M (unit).
-Definition _update_SV32_PTE_PPNi (v : SV32_PTE) (x : mword 22)
-: SV32_PTE :=
-
+Definition _update_SV32_PTE_PPNi (v : SV32_PTE) (x : mword 22) : SV32_PTE :=
{[ v with
SV32_PTE_SV32_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV32_PTE_SV32_PTE_chunk_0) 31 10 (subrange_vec_dec x 21 0)) ]}.
-Definition _get_SV32_PTE_RSW (v : SV32_PTE)
-: mword 2 :=
-
+Definition _get_SV32_PTE_RSW (v : SV32_PTE) : mword 2 :=
subrange_vec_dec v.(SV32_PTE_SV32_PTE_chunk_0) 9 8.
-Definition _set_SV32_PTE_RSW (r_ref : register_ref regstate register_value SV32_PTE) (v : mword 2)
+Definition _set_SV32_PTE_RSW (r_ref : register_ref regstate register_value SV32_PTE) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12720,21 +12388,16 @@ Definition _set_SV32_PTE_RSW (r_ref : register_ref regstate register_value SV32_
write_reg r_ref r
: M (unit).
-Definition _update_SV32_PTE_RSW (v : SV32_PTE) (x : mword 2)
-: SV32_PTE :=
-
+Definition _update_SV32_PTE_RSW (v : SV32_PTE) (x : mword 2) : SV32_PTE :=
{[ v with
SV32_PTE_SV32_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV32_PTE_SV32_PTE_chunk_0) 9 8 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_SV32_PTE_BITS (v : SV32_PTE)
-: mword 8 :=
-
+Definition _get_SV32_PTE_BITS (v : SV32_PTE) : mword 8 :=
subrange_vec_dec v.(SV32_PTE_SV32_PTE_chunk_0) 7 0.
-Definition _set_SV32_PTE_BITS (r_ref : register_ref regstate register_value SV32_PTE) (v : mword 8)
+Definition _set_SV32_PTE_BITS (r_ref : register_ref regstate register_value SV32_PTE) (v : mword 8)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12744,22 +12407,16 @@ Definition _set_SV32_PTE_BITS (r_ref : register_ref regstate register_value SV32
write_reg r_ref r
: M (unit).
-Definition _update_SV32_PTE_BITS (v : SV32_PTE) (x : mword 8)
-: SV32_PTE :=
-
+Definition _update_SV32_PTE_BITS (v : SV32_PTE) (x : mword 8) : SV32_PTE :=
{[ v with
SV32_PTE_SV32_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV32_PTE_SV32_PTE_chunk_0) 7 0 (subrange_vec_dec x 7 0)) ]}.
-Definition curAsid64 (satp : mword 64)
-: mword 16 :=
-
+Definition curAsid64 (satp : mword 64) : mword 16 :=
let s := Mk_Satp64 satp in
_get_Satp64_Asid s.
-Definition curPTB64 (satp : mword 64)
-: mword 56 :=
-
+Definition curPTB64 (satp : mword 64) : mword 56 :=
let s := Mk_Satp64 satp in
shiftl (EXTZ 56 (_get_Satp64_PPN s)) PAGESIZE_BITS.
@@ -12771,20 +12428,15 @@ Definition PTE39_LOG_SIZE := 3.
Hint Unfold PTE39_LOG_SIZE : sail.
Definition PTE39_SIZE := 8.
Hint Unfold PTE39_SIZE : sail.
-Definition Mk_SV39_Vaddr (v : mword 39)
-: SV39_Vaddr :=
-
+Definition Mk_SV39_Vaddr (v : mword 39) : SV39_Vaddr :=
{| SV39_Vaddr_SV39_Vaddr_chunk_0 := (subrange_vec_dec v 38 0) |}.
-Definition _get_SV39_Vaddr_bits (v : SV39_Vaddr)
-: mword 39 :=
-
+Definition _get_SV39_Vaddr_bits (v : SV39_Vaddr) : mword 39 :=
subrange_vec_dec v.(SV39_Vaddr_SV39_Vaddr_chunk_0) 38 0.
Definition _set_SV39_Vaddr_bits
-(r_ref : register_ref regstate register_value SV39_Vaddr) (v : mword 39)
+(r_ref : register_ref regstate register_value SV39_Vaddr) (v : mword 39)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12794,22 +12446,17 @@ Definition _set_SV39_Vaddr_bits
write_reg r_ref r
: M (unit).
-Definition _update_SV39_Vaddr_bits (v : SV39_Vaddr) (x : mword 39)
-: SV39_Vaddr :=
-
+Definition _update_SV39_Vaddr_bits (v : SV39_Vaddr) (x : mword 39) : SV39_Vaddr :=
{[ v with
SV39_Vaddr_SV39_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV39_Vaddr_SV39_Vaddr_chunk_0) 38 0 (subrange_vec_dec x 38 0)) ]}.
-Definition _get_SV39_Vaddr_VPNi (v : SV39_Vaddr)
-: mword 27 :=
-
+Definition _get_SV39_Vaddr_VPNi (v : SV39_Vaddr) : mword 27 :=
subrange_vec_dec v.(SV39_Vaddr_SV39_Vaddr_chunk_0) 38 12.
Definition _set_SV39_Vaddr_VPNi
-(r_ref : register_ref regstate register_value SV39_Vaddr) (v : mword 27)
+(r_ref : register_ref regstate register_value SV39_Vaddr) (v : mword 27)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12819,22 +12466,17 @@ Definition _set_SV39_Vaddr_VPNi
write_reg r_ref r
: M (unit).
-Definition _update_SV39_Vaddr_VPNi (v : SV39_Vaddr) (x : mword 27)
-: SV39_Vaddr :=
-
+Definition _update_SV39_Vaddr_VPNi (v : SV39_Vaddr) (x : mword 27) : SV39_Vaddr :=
{[ v with
SV39_Vaddr_SV39_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV39_Vaddr_SV39_Vaddr_chunk_0) 38 12 (subrange_vec_dec x 26 0)) ]}.
-Definition _get_SV39_Vaddr_PgOfs (v : SV39_Vaddr)
-: mword 12 :=
-
+Definition _get_SV39_Vaddr_PgOfs (v : SV39_Vaddr) : mword 12 :=
subrange_vec_dec v.(SV39_Vaddr_SV39_Vaddr_chunk_0) 11 0.
Definition _set_SV39_Vaddr_PgOfs
-(r_ref : register_ref regstate register_value SV39_Vaddr) (v : mword 12)
+(r_ref : register_ref regstate register_value SV39_Vaddr) (v : mword 12)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12844,27 +12486,20 @@ Definition _set_SV39_Vaddr_PgOfs
write_reg r_ref r
: M (unit).
-Definition _update_SV39_Vaddr_PgOfs (v : SV39_Vaddr) (x : mword 12)
-: SV39_Vaddr :=
-
+Definition _update_SV39_Vaddr_PgOfs (v : SV39_Vaddr) (x : mword 12) : SV39_Vaddr :=
{[ v with
SV39_Vaddr_SV39_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV39_Vaddr_SV39_Vaddr_chunk_0) 11 0 (subrange_vec_dec x 11 0)) ]}.
-Definition Mk_SV39_Paddr (v : mword 56)
-: SV39_Paddr :=
-
+Definition Mk_SV39_Paddr (v : mword 56) : SV39_Paddr :=
{| SV39_Paddr_SV39_Paddr_chunk_0 := (subrange_vec_dec v 55 0) |}.
-Definition _get_SV39_Paddr_bits (v : SV39_Paddr)
-: mword 56 :=
-
+Definition _get_SV39_Paddr_bits (v : SV39_Paddr) : mword 56 :=
subrange_vec_dec v.(SV39_Paddr_SV39_Paddr_chunk_0) 55 0.
Definition _set_SV39_Paddr_bits
-(r_ref : register_ref regstate register_value SV39_Paddr) (v : mword 56)
+(r_ref : register_ref regstate register_value SV39_Paddr) (v : mword 56)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12874,22 +12509,17 @@ Definition _set_SV39_Paddr_bits
write_reg r_ref r
: M (unit).
-Definition _update_SV39_Paddr_bits (v : SV39_Paddr) (x : mword 56)
-: SV39_Paddr :=
-
+Definition _update_SV39_Paddr_bits (v : SV39_Paddr) (x : mword 56) : SV39_Paddr :=
{[ v with
SV39_Paddr_SV39_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV39_Paddr_SV39_Paddr_chunk_0) 55 0 (subrange_vec_dec x 55 0)) ]}.
-Definition _get_SV39_Paddr_PPNi (v : SV39_Paddr)
-: mword 44 :=
-
+Definition _get_SV39_Paddr_PPNi (v : SV39_Paddr) : mword 44 :=
subrange_vec_dec v.(SV39_Paddr_SV39_Paddr_chunk_0) 55 12.
Definition _set_SV39_Paddr_PPNi
-(r_ref : register_ref regstate register_value SV39_Paddr) (v : mword 44)
+(r_ref : register_ref regstate register_value SV39_Paddr) (v : mword 44)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12899,22 +12529,17 @@ Definition _set_SV39_Paddr_PPNi
write_reg r_ref r
: M (unit).
-Definition _update_SV39_Paddr_PPNi (v : SV39_Paddr) (x : mword 44)
-: SV39_Paddr :=
-
+Definition _update_SV39_Paddr_PPNi (v : SV39_Paddr) (x : mword 44) : SV39_Paddr :=
{[ v with
SV39_Paddr_SV39_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV39_Paddr_SV39_Paddr_chunk_0) 55 12 (subrange_vec_dec x 43 0)) ]}.
-Definition _get_SV39_Paddr_PgOfs (v : SV39_Paddr)
-: mword 12 :=
-
+Definition _get_SV39_Paddr_PgOfs (v : SV39_Paddr) : mword 12 :=
subrange_vec_dec v.(SV39_Paddr_SV39_Paddr_chunk_0) 11 0.
Definition _set_SV39_Paddr_PgOfs
-(r_ref : register_ref regstate register_value SV39_Paddr) (v : mword 12)
+(r_ref : register_ref regstate register_value SV39_Paddr) (v : mword 12)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12924,27 +12549,19 @@ Definition _set_SV39_Paddr_PgOfs
write_reg r_ref r
: M (unit).
-Definition _update_SV39_Paddr_PgOfs (v : SV39_Paddr) (x : mword 12)
-: SV39_Paddr :=
-
+Definition _update_SV39_Paddr_PgOfs (v : SV39_Paddr) (x : mword 12) : SV39_Paddr :=
{[ v with
SV39_Paddr_SV39_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV39_Paddr_SV39_Paddr_chunk_0) 11 0 (subrange_vec_dec x 11 0)) ]}.
-Definition Mk_SV39_PTE (v : mword 64)
-: SV39_PTE :=
-
+Definition Mk_SV39_PTE (v : mword 64) : SV39_PTE :=
{| SV39_PTE_SV39_PTE_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_SV39_PTE_bits (v : SV39_PTE)
-: mword 64 :=
-
+Definition _get_SV39_PTE_bits (v : SV39_PTE) : mword 64 :=
subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 63 0.
Definition _set_SV39_PTE_bits (r_ref : register_ref regstate register_value SV39_PTE) (v : mword 64)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12954,22 +12571,35 @@ Definition _set_SV39_PTE_bits (r_ref : register_ref regstate register_value SV39
write_reg r_ref r
: M (unit).
-Definition _update_SV39_PTE_bits (v : SV39_PTE) (x : mword 64)
-: SV39_PTE :=
-
+Definition _update_SV39_PTE_bits (v : SV39_PTE) (x : mword 64) : SV39_PTE :=
{[ v with
SV39_PTE_SV39_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_SV39_PTE_PPNi (v : SV39_PTE)
-: mword 44 :=
-
+Definition _get_SV39_PTE_Ext (v : SV39_PTE) : mword 10 :=
+ subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 63 54.
+
+Definition _set_SV39_PTE_Ext (r_ref : register_ref regstate register_value SV39_PTE) (v : mword 10)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ SV39_PTE_SV39_PTE_chunk_0 :=
+ (update_subrange_vec_dec r.(SV39_PTE_SV39_PTE_chunk_0) 63 54 (subrange_vec_dec v 9 0)) ]}
+ : SV39_PTE in
+ write_reg r_ref r
+ : M (unit).
+
+Definition _update_SV39_PTE_Ext (v : SV39_PTE) (x : mword 10) : SV39_PTE :=
+ {[ v with
+ SV39_PTE_SV39_PTE_chunk_0 :=
+ (update_subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 63 54 (subrange_vec_dec x 9 0)) ]}.
+
+Definition _get_SV39_PTE_PPNi (v : SV39_PTE) : mword 44 :=
subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 53 10.
Definition _set_SV39_PTE_PPNi (r_ref : register_ref regstate register_value SV39_PTE) (v : mword 44)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -12979,21 +12609,16 @@ Definition _set_SV39_PTE_PPNi (r_ref : register_ref regstate register_value SV39
write_reg r_ref r
: M (unit).
-Definition _update_SV39_PTE_PPNi (v : SV39_PTE) (x : mword 44)
-: SV39_PTE :=
-
+Definition _update_SV39_PTE_PPNi (v : SV39_PTE) (x : mword 44) : SV39_PTE :=
{[ v with
SV39_PTE_SV39_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 53 10 (subrange_vec_dec x 43 0)) ]}.
-Definition _get_SV39_PTE_RSW (v : SV39_PTE)
-: mword 2 :=
-
+Definition _get_SV39_PTE_RSW (v : SV39_PTE) : mword 2 :=
subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 9 8.
-Definition _set_SV39_PTE_RSW (r_ref : register_ref regstate register_value SV39_PTE) (v : mword 2)
+Definition _set_SV39_PTE_RSW (r_ref : register_ref regstate register_value SV39_PTE) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13003,21 +12628,16 @@ Definition _set_SV39_PTE_RSW (r_ref : register_ref regstate register_value SV39_
write_reg r_ref r
: M (unit).
-Definition _update_SV39_PTE_RSW (v : SV39_PTE) (x : mword 2)
-: SV39_PTE :=
-
+Definition _update_SV39_PTE_RSW (v : SV39_PTE) (x : mword 2) : SV39_PTE :=
{[ v with
SV39_PTE_SV39_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 9 8 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_SV39_PTE_BITS (v : SV39_PTE)
-: mword 8 :=
-
+Definition _get_SV39_PTE_BITS (v : SV39_PTE) : mword 8 :=
subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 7 0.
-Definition _set_SV39_PTE_BITS (r_ref : register_ref regstate register_value SV39_PTE) (v : mword 8)
+Definition _set_SV39_PTE_BITS (r_ref : register_ref regstate register_value SV39_PTE) (v : mword 8)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13027,9 +12647,7 @@ Definition _set_SV39_PTE_BITS (r_ref : register_ref regstate register_value SV39
write_reg r_ref r
: M (unit).
-Definition _update_SV39_PTE_BITS (v : SV39_PTE) (x : mword 8)
-: SV39_PTE :=
-
+Definition _update_SV39_PTE_BITS (v : SV39_PTE) (x : mword 8) : SV39_PTE :=
{[ v with
SV39_PTE_SV39_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV39_PTE_SV39_PTE_chunk_0) 7 0 (subrange_vec_dec x 7 0)) ]}.
@@ -13042,20 +12660,15 @@ Definition PTE48_LOG_SIZE := 3.
Hint Unfold PTE48_LOG_SIZE : sail.
Definition PTE48_SIZE := 8.
Hint Unfold PTE48_SIZE : sail.
-Definition Mk_SV48_Vaddr (v : mword 48)
-: SV48_Vaddr :=
-
+Definition Mk_SV48_Vaddr (v : mword 48) : SV48_Vaddr :=
{| SV48_Vaddr_SV48_Vaddr_chunk_0 := (subrange_vec_dec v 47 0) |}.
-Definition _get_SV48_Vaddr_bits (v : SV48_Vaddr)
-: mword 48 :=
-
+Definition _get_SV48_Vaddr_bits (v : SV48_Vaddr) : mword 48 :=
subrange_vec_dec v.(SV48_Vaddr_SV48_Vaddr_chunk_0) 47 0.
Definition _set_SV48_Vaddr_bits
-(r_ref : register_ref regstate register_value SV48_Vaddr) (v : mword 48)
+(r_ref : register_ref regstate register_value SV48_Vaddr) (v : mword 48)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13065,22 +12678,17 @@ Definition _set_SV48_Vaddr_bits
write_reg r_ref r
: M (unit).
-Definition _update_SV48_Vaddr_bits (v : SV48_Vaddr) (x : mword 48)
-: SV48_Vaddr :=
-
+Definition _update_SV48_Vaddr_bits (v : SV48_Vaddr) (x : mword 48) : SV48_Vaddr :=
{[ v with
SV48_Vaddr_SV48_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV48_Vaddr_SV48_Vaddr_chunk_0) 47 0 (subrange_vec_dec x 47 0)) ]}.
-Definition _get_SV48_Vaddr_VPNi (v : SV48_Vaddr)
-: mword 27 :=
-
+Definition _get_SV48_Vaddr_VPNi (v : SV48_Vaddr) : mword 27 :=
subrange_vec_dec v.(SV48_Vaddr_SV48_Vaddr_chunk_0) 38 12.
Definition _set_SV48_Vaddr_VPNi
-(r_ref : register_ref regstate register_value SV48_Vaddr) (v : mword 27)
+(r_ref : register_ref regstate register_value SV48_Vaddr) (v : mword 27)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13090,22 +12698,17 @@ Definition _set_SV48_Vaddr_VPNi
write_reg r_ref r
: M (unit).
-Definition _update_SV48_Vaddr_VPNi (v : SV48_Vaddr) (x : mword 27)
-: SV48_Vaddr :=
-
+Definition _update_SV48_Vaddr_VPNi (v : SV48_Vaddr) (x : mword 27) : SV48_Vaddr :=
{[ v with
SV48_Vaddr_SV48_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV48_Vaddr_SV48_Vaddr_chunk_0) 38 12 (subrange_vec_dec x 26 0)) ]}.
-Definition _get_SV48_Vaddr_PgOfs (v : SV48_Vaddr)
-: mword 12 :=
-
+Definition _get_SV48_Vaddr_PgOfs (v : SV48_Vaddr) : mword 12 :=
subrange_vec_dec v.(SV48_Vaddr_SV48_Vaddr_chunk_0) 11 0.
Definition _set_SV48_Vaddr_PgOfs
-(r_ref : register_ref regstate register_value SV48_Vaddr) (v : mword 12)
+(r_ref : register_ref regstate register_value SV48_Vaddr) (v : mword 12)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13115,27 +12718,20 @@ Definition _set_SV48_Vaddr_PgOfs
write_reg r_ref r
: M (unit).
-Definition _update_SV48_Vaddr_PgOfs (v : SV48_Vaddr) (x : mword 12)
-: SV48_Vaddr :=
-
+Definition _update_SV48_Vaddr_PgOfs (v : SV48_Vaddr) (x : mword 12) : SV48_Vaddr :=
{[ v with
SV48_Vaddr_SV48_Vaddr_chunk_0 :=
(update_subrange_vec_dec v.(SV48_Vaddr_SV48_Vaddr_chunk_0) 11 0 (subrange_vec_dec x 11 0)) ]}.
-Definition Mk_SV48_Paddr (v : mword 56)
-: SV48_Paddr :=
-
+Definition Mk_SV48_Paddr (v : mword 56) : SV48_Paddr :=
{| SV48_Paddr_SV48_Paddr_chunk_0 := (subrange_vec_dec v 55 0) |}.
-Definition _get_SV48_Paddr_bits (v : SV48_Paddr)
-: mword 56 :=
-
+Definition _get_SV48_Paddr_bits (v : SV48_Paddr) : mword 56 :=
subrange_vec_dec v.(SV48_Paddr_SV48_Paddr_chunk_0) 55 0.
Definition _set_SV48_Paddr_bits
-(r_ref : register_ref regstate register_value SV48_Paddr) (v : mword 56)
+(r_ref : register_ref regstate register_value SV48_Paddr) (v : mword 56)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13145,22 +12741,17 @@ Definition _set_SV48_Paddr_bits
write_reg r_ref r
: M (unit).
-Definition _update_SV48_Paddr_bits (v : SV48_Paddr) (x : mword 56)
-: SV48_Paddr :=
-
+Definition _update_SV48_Paddr_bits (v : SV48_Paddr) (x : mword 56) : SV48_Paddr :=
{[ v with
SV48_Paddr_SV48_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV48_Paddr_SV48_Paddr_chunk_0) 55 0 (subrange_vec_dec x 55 0)) ]}.
-Definition _get_SV48_Paddr_PPNi (v : SV48_Paddr)
-: mword 44 :=
-
+Definition _get_SV48_Paddr_PPNi (v : SV48_Paddr) : mword 44 :=
subrange_vec_dec v.(SV48_Paddr_SV48_Paddr_chunk_0) 55 12.
Definition _set_SV48_Paddr_PPNi
-(r_ref : register_ref regstate register_value SV48_Paddr) (v : mword 44)
+(r_ref : register_ref regstate register_value SV48_Paddr) (v : mword 44)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13170,22 +12761,17 @@ Definition _set_SV48_Paddr_PPNi
write_reg r_ref r
: M (unit).
-Definition _update_SV48_Paddr_PPNi (v : SV48_Paddr) (x : mword 44)
-: SV48_Paddr :=
-
+Definition _update_SV48_Paddr_PPNi (v : SV48_Paddr) (x : mword 44) : SV48_Paddr :=
{[ v with
SV48_Paddr_SV48_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV48_Paddr_SV48_Paddr_chunk_0) 55 12 (subrange_vec_dec x 43 0)) ]}.
-Definition _get_SV48_Paddr_PgOfs (v : SV48_Paddr)
-: mword 12 :=
-
+Definition _get_SV48_Paddr_PgOfs (v : SV48_Paddr) : mword 12 :=
subrange_vec_dec v.(SV48_Paddr_SV48_Paddr_chunk_0) 11 0.
Definition _set_SV48_Paddr_PgOfs
-(r_ref : register_ref regstate register_value SV48_Paddr) (v : mword 12)
+(r_ref : register_ref regstate register_value SV48_Paddr) (v : mword 12)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13195,27 +12781,19 @@ Definition _set_SV48_Paddr_PgOfs
write_reg r_ref r
: M (unit).
-Definition _update_SV48_Paddr_PgOfs (v : SV48_Paddr) (x : mword 12)
-: SV48_Paddr :=
-
+Definition _update_SV48_Paddr_PgOfs (v : SV48_Paddr) (x : mword 12) : SV48_Paddr :=
{[ v with
SV48_Paddr_SV48_Paddr_chunk_0 :=
(update_subrange_vec_dec v.(SV48_Paddr_SV48_Paddr_chunk_0) 11 0 (subrange_vec_dec x 11 0)) ]}.
-Definition Mk_SV48_PTE (v : mword 64)
-: SV48_PTE :=
-
+Definition Mk_SV48_PTE (v : mword 64) : SV48_PTE :=
{| SV48_PTE_SV48_PTE_chunk_0 := (subrange_vec_dec v 63 0) |}.
-Definition _get_SV48_PTE_bits (v : SV48_PTE)
-: mword 64 :=
-
+Definition _get_SV48_PTE_bits (v : SV48_PTE) : mword 64 :=
subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 63 0.
Definition _set_SV48_PTE_bits (r_ref : register_ref regstate register_value SV48_PTE) (v : mword 64)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13225,22 +12803,35 @@ Definition _set_SV48_PTE_bits (r_ref : register_ref regstate register_value SV48
write_reg r_ref r
: M (unit).
-Definition _update_SV48_PTE_bits (v : SV48_PTE) (x : mword 64)
-: SV48_PTE :=
-
+Definition _update_SV48_PTE_bits (v : SV48_PTE) (x : mword 64) : SV48_PTE :=
{[ v with
SV48_PTE_SV48_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 63 0 (subrange_vec_dec x 63 0)) ]}.
-Definition _get_SV48_PTE_PPNi (v : SV48_PTE)
-: mword 44 :=
-
+Definition _get_SV48_PTE_Ext (v : SV48_PTE) : mword 10 :=
+ subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 63 54.
+
+Definition _set_SV48_PTE_Ext (r_ref : register_ref regstate register_value SV48_PTE) (v : mword 10)
+: M (unit) :=
+ (reg_deref r_ref) >>= fun r =>
+ let r :=
+ {[ r with
+ SV48_PTE_SV48_PTE_chunk_0 :=
+ (update_subrange_vec_dec r.(SV48_PTE_SV48_PTE_chunk_0) 63 54 (subrange_vec_dec v 9 0)) ]}
+ : SV48_PTE in
+ write_reg r_ref r
+ : M (unit).
+
+Definition _update_SV48_PTE_Ext (v : SV48_PTE) (x : mword 10) : SV48_PTE :=
+ {[ v with
+ SV48_PTE_SV48_PTE_chunk_0 :=
+ (update_subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 63 54 (subrange_vec_dec x 9 0)) ]}.
+
+Definition _get_SV48_PTE_PPNi (v : SV48_PTE) : mword 44 :=
subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 53 10.
Definition _set_SV48_PTE_PPNi (r_ref : register_ref regstate register_value SV48_PTE) (v : mword 44)
-
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13250,21 +12841,16 @@ Definition _set_SV48_PTE_PPNi (r_ref : register_ref regstate register_value SV48
write_reg r_ref r
: M (unit).
-Definition _update_SV48_PTE_PPNi (v : SV48_PTE) (x : mword 44)
-: SV48_PTE :=
-
+Definition _update_SV48_PTE_PPNi (v : SV48_PTE) (x : mword 44) : SV48_PTE :=
{[ v with
SV48_PTE_SV48_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 53 10 (subrange_vec_dec x 43 0)) ]}.
-Definition _get_SV48_PTE_RSW (v : SV48_PTE)
-: mword 2 :=
-
+Definition _get_SV48_PTE_RSW (v : SV48_PTE) : mword 2 :=
subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 9 8.
-Definition _set_SV48_PTE_RSW (r_ref : register_ref regstate register_value SV48_PTE) (v : mword 2)
+Definition _set_SV48_PTE_RSW (r_ref : register_ref regstate register_value SV48_PTE) (v : mword 2)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13274,21 +12860,16 @@ Definition _set_SV48_PTE_RSW (r_ref : register_ref regstate register_value SV48_
write_reg r_ref r
: M (unit).
-Definition _update_SV48_PTE_RSW (v : SV48_PTE) (x : mword 2)
-: SV48_PTE :=
-
+Definition _update_SV48_PTE_RSW (v : SV48_PTE) (x : mword 2) : SV48_PTE :=
{[ v with
SV48_PTE_SV48_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 9 8 (subrange_vec_dec x 1 0)) ]}.
-Definition _get_SV48_PTE_BITS (v : SV48_PTE)
-: mword 8 :=
-
+Definition _get_SV48_PTE_BITS (v : SV48_PTE) : mword 8 :=
subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 7 0.
-Definition _set_SV48_PTE_BITS (r_ref : register_ref regstate register_value SV48_PTE) (v : mword 8)
+Definition _set_SV48_PTE_BITS (r_ref : register_ref regstate register_value SV48_PTE) (v : mword 8)
: M (unit) :=
-
(reg_deref r_ref) >>= fun r =>
let r :=
{[ r with
@@ -13298,66 +12879,59 @@ Definition _set_SV48_PTE_BITS (r_ref : register_ref regstate register_value SV48
write_reg r_ref r
: M (unit).
-Definition _update_SV48_PTE_BITS (v : SV48_PTE) (x : mword 8)
-: SV48_PTE :=
-
+Definition _update_SV48_PTE_BITS (v : SV48_PTE) (x : mword 8) : SV48_PTE :=
{[ v with
SV48_PTE_SV48_PTE_chunk_0 :=
(update_subrange_vec_dec v.(SV48_PTE_SV48_PTE_chunk_0) 7 0 (subrange_vec_dec x 7 0)) ]}.
Definition make_TLB_Entry {asidlen : Z} {valen : Z} {palen : Z} {ptelen : Z}
(asid : mword asidlen) (global : bool) (vAddr : mword valen) (pAddr : mword palen)
-(pte : mword ptelen) (level : Z) (pteAddr : mword palen) (levelBitSize : Z) `{ArithFact (valen > 0)}
-`{ArithFact (0 <= level)} `{ArithFact (0 <= levelBitSize)}
+(pte : mword ptelen) (level : Z) (pteAddr : mword palen) (levelBitSize : Z)
+`{ArithFact (valen >? 0)} `{ArithFact (0 <=? level)} `{ArithFact (0 <=? levelBitSize)}
: M (TLB_Entry asidlen valen palen ptelen) :=
-
let shift := Z.add PAGESIZE_BITS (Z.mul level levelBitSize) in
let vAddrMask : bits valen :=
sub_vec_int
- (shiftl
- (xor_vec vAddr (xor_vec vAddr (EXTZ (length_mword vAddr) (vec_of_bits [B1] : mword 1))))
- shift) 1 in
+ (shiftl (xor_vec vAddr (xor_vec vAddr (EXTZ (length_mword vAddr) ('b"1" : mword 1)))) shift)
+ 1 in
let vMatchMask : bits valen := not_vec vAddrMask in
((read_reg mcycle_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- returnm (({| TLB_Entry_asid := asid;
- TLB_Entry_global := global;
- TLB_Entry_pte := pte;
- TLB_Entry_pteAddr := pteAddr;
- TLB_Entry_vAddrMask := vAddrMask;
- TLB_Entry_vMatchMask := vMatchMask;
- TLB_Entry_vAddr := (and_vec vAddr vMatchMask);
- TLB_Entry_pAddr := (shiftl (shiftr pAddr shift) shift);
- TLB_Entry_age := w__0 |})
- : TLB_Entry asidlen valen palen ptelen).
+ returnm ({| TLB_Entry_asid := asid;
+ TLB_Entry_global := global;
+ TLB_Entry_pte := pte;
+ TLB_Entry_pteAddr := pteAddr;
+ TLB_Entry_vAddrMask := vAddrMask;
+ TLB_Entry_vMatchMask := vMatchMask;
+ TLB_Entry_vAddr := (and_vec vAddr vMatchMask);
+ TLB_Entry_pAddr := (shiftl (shiftr pAddr shift) shift);
+ TLB_Entry_age := w__0 |}).
Definition match_TLB_Entry {asidlen : Z} {valen : Z} {palen : Z} {ptelen : Z}
-(ent : TLB_Entry asidlen valen palen ptelen) (asid : mword asidlen) (vaddr : mword valen)
+(ent : TLB_Entry asidlen valen palen ptelen) (asid : mword asidlen) (vaddr : mword valen)
: bool :=
-
andb (orb ent.(TLB_Entry_global) (eq_vec ent.(TLB_Entry_asid) asid))
(eq_vec ent.(TLB_Entry_vAddr) (and_vec ent.(TLB_Entry_vMatchMask) vaddr)).
Definition flush_TLB_Entry {asidlen : Z} {valen : Z} {palen : Z} {ptelen : Z}
(e : TLB_Entry asidlen valen palen ptelen) (asid : option (mword asidlen))
-(addr : option (mword valen))
+(addr : option (mword valen))
: bool :=
-
match (asid, addr) with
| (None, None) => true
- | (None, Some (a)) => eq_vec e.(TLB_Entry_vAddr) (and_vec e.(TLB_Entry_vMatchMask) a)
- | (Some (i), None) => andb (eq_vec e.(TLB_Entry_asid) i) (negb e.(TLB_Entry_global))
- | (Some (i), Some (a)) =>
+ | (None, Some a) => eq_vec e.(TLB_Entry_vAddr) (and_vec e.(TLB_Entry_vMatchMask) a)
+ | (Some i, None) => andb (eq_vec e.(TLB_Entry_asid) i) (negb e.(TLB_Entry_global))
+ | (Some i, Some a) =>
andb (eq_vec e.(TLB_Entry_asid) i)
(andb (eq_vec e.(TLB_Entry_vAddr) (and_vec a e.(TLB_Entry_vMatchMask)))
(negb e.(TLB_Entry_global)))
end.
Fixpoint _rec_walk39
-(vaddr : mword 39) (ac : AccessType) (priv : Privilege) (mxr : bool) (do_sum : bool)
-(ptb : mword 56) (level : Z) (global : bool) (_reclimit : Z) `{ArithFact (0 <= level)}
-(_acc : Acc (Zwf 0) _reclimit)
-{struct _acc} : M (PTW_Result (mword 56) SV39_PTE) :=
-
+(vaddr : mword 39) (ac : AccessType unit) (priv : Privilege) (mxr : bool) (do_sum : bool)
+(ptb : mword 56) (level : Z) (global : bool) (ext_ptw : unit) (_reclimit : Z)
+`{ArithFact (0 <=? level)} (_acc : Acc (Zwf 0) _reclimit)
+{struct _acc} : M (PTW_Result (mword 56) SV39_PTE).
+exact (
assert_exp' (Z.geb _reclimit 0) "recursion limit reached" >>= fun _ =>
let va := Mk_SV39_Vaddr vaddr in
let pt_ofs : paddr64 :=
@@ -13368,180 +12942,159 @@ Fixpoint _rec_walk39
let pte_addr := add_vec ptb pt_ofs in
(mem_read ac (EXTZ 64 pte_addr) 8 false false false) >>= fun w__0 : MemoryOpResult (mword (8 * 8)) =>
(match w__0 with
- | MemException (_) => returnm ((PTW_Failure (PTW_Access)) : PTW_Result (mword 56) SV39_PTE)
- | MemValue (v) =>
+ | MemException _ => returnm (PTW_Failure (PTW_Access tt, ext_ptw))
+ | MemValue v =>
let pte := Mk_SV39_PTE v in
let pbits := _get_SV39_PTE_BITS pte in
+ let ext_pte := _get_SV39_PTE_Ext pte in
let pattr := Mk_PTE_Bits pbits in
- let is_global := orb global (eq_vec (_get_PTE_Bits_G pattr) ((bool_to_bits true) : mword 1)) in
- (if ((isInvalidPTE pbits)) then
- returnm ((PTW_Failure
- (PTW_Invalid_PTE))
- : PTW_Result (mword 56) SV39_PTE)
- else if ((isPTEPtr pbits)) then
- (if sumbool_of_bool ((Z.eqb level 0)) then
- returnm ((PTW_Failure
- (PTW_Invalid_PTE))
- : PTW_Result (mword 56) SV39_PTE)
- else
+ let is_global := orb global (eq_vec (_get_PTE_Bits_G pattr) ('b"1" : mword 1)) in
+ (if isInvalidPTE pbits ext_pte then returnm (PTW_Failure (PTW_Invalid_PTE tt, ext_ptw))
+ else if isPTEPtr pbits ext_pte then
+ (if sumbool_of_bool (Z.gtb level 0) then
(_rec_walk39 vaddr ac priv mxr do_sum
(shiftl (EXTZ 56 (_get_SV39_PTE_PPNi pte)) PAGESIZE_BITS) (Z.sub level 1) is_global
- (Z.sub _reclimit 1) (_limit_reduces _acc))
- : M (PTW_Result (mword 56) SV39_PTE))
+ ext_ptw (Z.sub _reclimit 1) _ (_limit_reduces _acc))
+ : M (PTW_Result (mword 56) SV39_PTE)
+ else returnm (PTW_Failure (PTW_Invalid_PTE tt, ext_ptw)))
: M (PTW_Result (mword 56) SV39_PTE)
else
- (checkPTEPermission ac priv mxr do_sum pattr) >>= fun w__3 : bool =>
- returnm ((if sumbool_of_bool ((negb w__3)) then PTW_Failure (PTW_No_Permission)
- else if sumbool_of_bool ((Z.gtb level 0)) then
- let mask :=
- sub_vec_int
- (shiftl
- (xor_vec (_get_SV39_PTE_PPNi pte)
- (xor_vec (_get_SV39_PTE_PPNi pte)
- (EXTZ 44 (vec_of_bits [B1] : mword 1))))
- (Z.mul level SV39_LEVEL_BITS)) 1 in
- if ((neq_vec (and_vec (_get_SV39_PTE_PPNi pte) mask)
- (EXTZ 44 (vec_of_bits [B0] : mword 1)))) then
- PTW_Failure
- (PTW_Misaligned)
+ (checkPTEPermission ac priv mxr do_sum pattr ext_pte ext_ptw) >>= fun w__3 : PTE_Check =>
+ returnm (match w__3 with
+ | PTE_Check_Failure ext_ptw => PTW_Failure (PTW_No_Permission tt, ext_ptw)
+ | PTE_Check_Success ext_ptw =>
+ if sumbool_of_bool (Z.gtb level 0) then
+ let mask :=
+ sub_vec_int
+ (shiftl
+ (xor_vec (_get_SV39_PTE_PPNi pte)
+ (xor_vec (_get_SV39_PTE_PPNi pte) (EXTZ 44 ('b"1" : mword 1))))
+ (Z.mul level SV39_LEVEL_BITS)) 1 in
+ if neq_vec (and_vec (_get_SV39_PTE_PPNi pte) mask)
+ (EXTZ 44 ('b"0" : mword 1)) then
+ PTW_Failure (PTW_Misaligned tt, ext_ptw)
+ else
+ let ppn :=
+ or_vec (_get_SV39_PTE_PPNi pte)
+ (and_vec (EXTZ 44 (_get_SV39_Vaddr_VPNi va)) mask) in
+ PTW_Success
+ (concat_vec ppn (_get_SV39_Vaddr_PgOfs va), pte, pte_addr, build_ex
+ level, is_global, ext_ptw)
else
- let ppn :=
- or_vec (_get_SV39_PTE_PPNi pte)
- (and_vec (EXTZ 44 (_get_SV39_Vaddr_VPNi va)) mask) in
PTW_Success
- ((concat_vec ppn (_get_SV39_Vaddr_PgOfs va), pte, pte_addr, build_ex
- level, is_global))
- else
- PTW_Success
- ((concat_vec (_get_SV39_PTE_PPNi pte) (_get_SV39_Vaddr_PgOfs va), pte, pte_addr, build_ex
- level, is_global)))
- : PTW_Result (mword 56) SV39_PTE))
+ (concat_vec (_get_SV39_PTE_PPNi pte) (_get_SV39_Vaddr_PgOfs va), pte, pte_addr, build_ex
+ level, is_global, ext_ptw)
+ end))
: M (PTW_Result (mword 56) SV39_PTE)
end)
- : M (PTW_Result (mword 56) SV39_PTE).
+ : M (PTW_Result (mword 56) SV39_PTE)
+).
+Defined.
+
Definition walk39
-(_arg0 : mword 39) (_arg1 : AccessType) (_arg2 : Privilege) (_arg3 : bool) (_arg4 : bool)
-(_arg5 : mword 56) (level : Z) (_arg7 : bool) `{ArithFact (0 <= level)}
+(_arg0 : mword 39) (_arg1 : AccessType unit) (_arg2 : Privilege) (_arg3 : bool) (_arg4 : bool)
+(_arg5 : mword 56) (level : Z) (_arg7 : bool) (_arg8 : unit) `{ArithFact (0 <=? level)}
: M (PTW_Result (mword 56) SV39_PTE) :=
-
- (_rec_walk39 _arg0 _arg1 _arg2 _arg3 _arg4 _arg5 level _arg7 (level : Z) (Zwf_guarded _))
+ (_rec_walk39 _arg0 _arg1 _arg2 _arg3 _arg4 _arg5 level _arg7 _arg8 (level : Z) (Zwf_guarded _))
: M (PTW_Result (mword 56) SV39_PTE).
-Definition lookup_TLB39 (asid : mword 16) (vaddr : mword 39)
-: M (option (({n : Z & ArithFact (n >= 0)} * TLB_Entry 16 39 56 64))) :=
-
+Definition lookup_TLB39 (asid : mword 16) (vaddr : mword 39)
+: M (option (({n : Z & ArithFact (n >=? 0)} * TLB_Entry 16 39 56 64))) :=
read_reg tlb39_ref >>= fun w__0 : option (TLB_Entry 16 39 56 64) =>
- returnm ((match w__0 with
- | None => None
- | Some (e) => if ((match_TLB_Entry e asid vaddr)) then Some ((build_ex 0, e)) else None
- end)
- : option (({n : Z & ArithFact (n >= 0)} * TLB_Entry 16 39 56 64))).
+ returnm (match w__0 with
+ | None => None
+ | Some e => if match_TLB_Entry e asid vaddr then Some (build_ex 0, e) else None
+ end).
Definition add_to_TLB39
(asid : mword 16) (vAddr : mword 39) (pAddr : mword 56) (pte : SV39_PTE) (pteAddr : mword 56)
-(level : Z) (global : bool) `{ArithFact (0 <= level)}
+(level : Z) (global : bool) `{ArithFact (0 <=? level)}
: M (unit) :=
-
(make_TLB_Entry asid global vAddr pAddr (_get_SV39_PTE_bits pte) level pteAddr SV39_LEVEL_BITS) >>= fun ent : TLB39_Entry =>
- write_reg tlb39_ref (Some (ent))
+ write_reg tlb39_ref (Some ent)
: M (unit).
-Definition write_TLB39 (idx : Z) (ent : TLB_Entry 16 39 56 64) `{ArithFact (0 <= idx)}
-: M (unit) :=
-
- write_reg tlb39_ref (Some (ent))
- : M (unit).
+Definition write_TLB39 (idx : Z) (ent : TLB_Entry 16 39 56 64) `{ArithFact (0 <=? idx)} : M (unit) :=
+ write_reg tlb39_ref (Some ent) : M (unit).
-Definition flush_TLB39 (asid : option (mword 16)) (addr : option (mword 39))
-: M (unit) :=
-
+Definition flush_TLB39 (asid : option (mword 16)) (addr : option (mword 39)) : M (unit) :=
read_reg tlb39_ref >>= fun w__0 : option (TLB_Entry 16 39 56 64) =>
(match w__0 with
- | None => returnm (tt : unit)
- | Some (e) =>
- (if ((flush_TLB_Entry e asid addr)) then write_reg tlb39_ref None : M (unit)
- else returnm (tt : unit))
+ | None => returnm tt
+ | Some e =>
+ (if flush_TLB_Entry e asid addr then write_reg tlb39_ref None : M (unit)
+ else returnm tt)
: M (unit)
end)
: M (unit).
Definition translate39
-(asid : mword 16) (ptb : mword 56) (vAddr : mword 39) (ac : AccessType) (priv : Privilege)
-(mxr : bool) (do_sum : bool) (level : Z) `{ArithFact (0 <= level)}
+(asid : mword 16) (ptb : mword 56) (vAddr : mword 39) (ac : AccessType unit) (priv : Privilege)
+(mxr : bool) (do_sum : bool) (level : Z) (ext_ptw : unit) `{ArithFact (0 <=? level)}
: M (TR_Result (mword 56) PTW_Error) :=
-
- (lookup_TLB39 asid vAddr) >>= fun w__0 : option (({n : Z & ArithFact (n >= 0)} * TLB_Entry 16 39 56 64)) =>
+ (lookup_TLB39 asid vAddr) >>= fun w__0 : option (({n : Z & ArithFact (n >=? 0)} * TLB_Entry 16 39 56 64)) =>
(match w__0 with
- | Some ((existT _ idx _, ent)) =>
+ | Some (existT _ idx _, ent) =>
let pte := Mk_SV39_PTE ent.(TLB_Entry_pte) in
+ let ext_pte := _get_SV39_PTE_Ext pte in
let pteBits := Mk_PTE_Bits (_get_SV39_PTE_BITS pte) in
- (checkPTEPermission ac priv mxr do_sum pteBits) >>= fun w__1 : bool =>
- (if sumbool_of_bool ((negb w__1)) then
- returnm ((TR_Failure
- (PTW_No_Permission))
- : TR_Result (mword 56) PTW_Error)
- else
- (match (update_PTE_Bits pteBits ac) with
- | None =>
- returnm ((TR_Address
+ (checkPTEPermission ac priv mxr do_sum pteBits ext_pte ext_ptw) >>= fun w__1 : PTE_Check =>
+ (match w__1 with
+ | PTE_Check_Failure ext_ptw => returnm (TR_Failure (PTW_No_Permission tt, ext_ptw))
+ | PTE_Check_Success ext_ptw =>
+ (match (update_PTE_Bits pteBits ac ext_pte) with
+ | None =>
+ returnm (TR_Address
(or_vec ent.(TLB_Entry_pAddr)
- (EXTZ 56 (and_vec vAddr ent.(TLB_Entry_vAddrMask)))))
- : TR_Result (mword 56) PTW_Error)
- | Some (pbits) =>
- (if ((negb (plat_enable_dirty_update tt))) then
- returnm ((TR_Failure
- (PTW_PTE_Update))
- : TR_Result (mword 56) PTW_Error)
- else
- let n_pte := _update_SV39_PTE_BITS pte (_get_PTE_Bits_bits pbits) in
- let n_ent : TLB39_Entry := ent in
- let n_ent :=
- {[ n_ent with TLB_Entry_pte := (_get_SV39_PTE_bits n_pte) ]}
- : TLB_Entry 16 39 56 64 in
- (write_TLB39 idx n_ent) >>
- (mem_write_value (EXTZ 64 ent.(TLB_Entry_pteAddr)) 8 (_get_SV39_PTE_bits n_pte)
- false false false) >>= fun w__2 : MemoryOpResult bool =>
- (match w__2 with
- | MemValue (_) => returnm (tt : unit)
- | MemException (e) =>
- (internal_error "invalid physical address in TLB") : M (unit)
- end) >>
- returnm ((TR_Address
+ (EXTZ 56 (and_vec vAddr ent.(TLB_Entry_vAddrMask))), ext_ptw))
+ | Some (pbits, ext) =>
+ (if negb (plat_enable_dirty_update tt) then
+ returnm (TR_Failure (PTW_PTE_Update tt, ext_ptw))
+ else
+ let n_pte := _update_SV39_PTE_BITS pte (_get_PTE_Bits_bits pbits) in
+ let n_pte := (_update_SV39_PTE_Ext n_pte ext) : SV39_PTE in
+ let n_ent : TLB39_Entry := ent in
+ let n_ent :=
+ {[ n_ent with TLB_Entry_pte := (_get_SV39_PTE_bits n_pte) ]}
+ : TLB_Entry 16 39 56 64 in
+ (write_TLB39 idx n_ent) >>
+ (mem_write_value (EXTZ 64 ent.(TLB_Entry_pteAddr)) 8 (_get_SV39_PTE_bits n_pte)
+ false false false) >>= fun w__2 : MemoryOpResult bool =>
+ (match w__2 with
+ | MemValue _ => returnm tt
+ | MemException e =>
+ (internal_error "invalid physical address in TLB") : M (unit)
+ end) >>
+ returnm (TR_Address
(or_vec ent.(TLB_Entry_pAddr)
- (EXTZ 56 (and_vec vAddr ent.(TLB_Entry_vAddrMask)))))
- : TR_Result (mword 56) PTW_Error))
- : M (TR_Result (mword 56) PTW_Error)
- end)
- : M (TR_Result (mword 56) PTW_Error))
+ (EXTZ 56 (and_vec vAddr ent.(TLB_Entry_vAddrMask))), ext_ptw)))
+ : M (TR_Result (mword 56) PTW_Error)
+ end)
+ : M (TR_Result (mword 56) PTW_Error)
+ end)
: M (TR_Result (mword 56) PTW_Error)
| None =>
- (walk39 vAddr ac priv mxr do_sum ptb level false) >>= fun w__6 : PTW_Result (mword 56) SV39_PTE =>
+ (walk39 vAddr ac priv mxr do_sum ptb level false ext_ptw) >>= fun w__6 : PTW_Result (mword 56) SV39_PTE =>
(match w__6 with
- | PTW_Failure (f) => returnm ((TR_Failure (f)) : TR_Result (mword 56) PTW_Error)
- | PTW_Success ((pAddr, pte, pteAddr, existT _ level _, global)) =>
- (match (update_PTE_Bits (Mk_PTE_Bits (_get_SV39_PTE_BITS pte)) ac) with
+ | PTW_Failure (f, ext_ptw) => returnm (TR_Failure (f, ext_ptw))
+ | PTW_Success (pAddr, pte, pteAddr, existT _ level _, global, ext_ptw) =>
+ (match (update_PTE_Bits (Mk_PTE_Bits (_get_SV39_PTE_BITS pte)) ac (_get_SV39_PTE_Ext pte)) with
| None =>
(add_to_TLB39 asid vAddr pAddr pte pteAddr level global) >>
- returnm ((TR_Address
- (pAddr))
- : TR_Result (mword 56) PTW_Error)
- | Some (pbits) =>
- (if ((negb (plat_enable_dirty_update tt))) then
- returnm ((TR_Failure
- (PTW_PTE_Update))
- : TR_Result (mword 56) PTW_Error)
+ returnm (TR_Address (pAddr, ext_ptw))
+ | Some (pbits, ext) =>
+ (if negb (plat_enable_dirty_update tt) then
+ returnm (TR_Failure (PTW_PTE_Update tt, ext_ptw))
else
let w_pte : SV39_PTE := _update_SV39_PTE_BITS pte (_get_PTE_Bits_bits pbits) in
+ let w_pte : SV39_PTE := (_update_SV39_PTE_Ext w_pte ext) : SV39_PTE in
(mem_write_value (EXTZ 64 pteAddr) 8 (_get_SV39_PTE_bits w_pte) false false false) >>= fun w__7 : MemoryOpResult bool =>
(match w__7 with
- | MemValue (_) =>
+ | MemValue _ =>
(add_to_TLB39 asid vAddr pAddr w_pte pteAddr level global) >>
- returnm ((TR_Address
- (pAddr))
- : TR_Result (mword 56) PTW_Error)
- | MemException (e) =>
- returnm ((TR_Failure (PTW_Access)) : TR_Result (mword 56) PTW_Error)
+ returnm (TR_Address (pAddr, ext_ptw))
+ | MemException e => returnm (TR_Failure (PTW_Access tt, ext_ptw))
end)
: M (TR_Result (mword 56) PTW_Error))
: M (TR_Result (mword 56) PTW_Error)
@@ -13552,14 +13105,14 @@ Definition translate39
end)
: M (TR_Result (mword 56) PTW_Error).
-Definition init_vmem_sv39 '(tt : unit) : M (unit) := write_reg tlb39_ref None : M (unit).
+Definition init_vmem_sv39 '(tt : unit) : M (unit) := write_reg tlb39_ref None : M (unit).
Fixpoint _rec_walk48
-(vaddr : mword 48) (ac : AccessType) (priv : Privilege) (mxr : bool) (do_sum : bool)
-(ptb : mword 56) (level : Z) (global : bool) (_reclimit : Z) `{ArithFact (0 <= level)}
-(_acc : Acc (Zwf 0) _reclimit)
-{struct _acc} : M (PTW_Result (mword 56) SV48_PTE) :=
-
+(vaddr : mword 48) (ac : AccessType unit) (priv : Privilege) (mxr : bool) (do_sum : bool)
+(ptb : mword 56) (level : Z) (global : bool) (ext_ptw : unit) (_reclimit : Z)
+`{ArithFact (0 <=? level)} (_acc : Acc (Zwf 0) _reclimit)
+{struct _acc} : M (PTW_Result (mword 56) SV48_PTE).
+exact (
assert_exp' (Z.geb _reclimit 0) "recursion limit reached" >>= fun _ =>
let va := Mk_SV48_Vaddr vaddr in
let pt_ofs : paddr64 :=
@@ -13570,135 +13123,117 @@ Fixpoint _rec_walk48
let pte_addr := add_vec ptb pt_ofs in
(mem_read ac (EXTZ 64 pte_addr) 8 false false false) >>= fun w__0 : MemoryOpResult (mword (8 * 8)) =>
(match w__0 with
- | MemException (_) => returnm ((PTW_Failure (PTW_Access)) : PTW_Result (mword 56) SV48_PTE)
- | MemValue (v) =>
+ | MemException _ => returnm (PTW_Failure (PTW_Access tt, ext_ptw))
+ | MemValue v =>
let pte := Mk_SV48_PTE v in
let pbits := _get_SV48_PTE_BITS pte in
+ let ext_pte := _get_SV48_PTE_Ext pte in
let pattr := Mk_PTE_Bits pbits in
- let is_global := orb global (eq_vec (_get_PTE_Bits_G pattr) ((bool_to_bits true) : mword 1)) in
- (if ((isInvalidPTE pbits)) then
- returnm ((PTW_Failure
- (PTW_Invalid_PTE))
- : PTW_Result (mword 56) SV48_PTE)
- else if ((isPTEPtr pbits)) then
- (if sumbool_of_bool ((Z.eqb level 0)) then
- returnm ((PTW_Failure
- (PTW_Invalid_PTE))
- : PTW_Result (mword 56) SV48_PTE)
- else
+ let is_global := orb global (eq_vec (_get_PTE_Bits_G pattr) ('b"1" : mword 1)) in
+ (if isInvalidPTE pbits ext_pte then returnm (PTW_Failure (PTW_Invalid_PTE tt, ext_ptw))
+ else if isPTEPtr pbits ext_pte then
+ (if sumbool_of_bool (Z.gtb level 0) then
(_rec_walk48 vaddr ac priv mxr do_sum
(shiftl (EXTZ 56 (_get_SV48_PTE_PPNi pte)) PAGESIZE_BITS) (Z.sub level 1) is_global
- (Z.sub _reclimit 1) (_limit_reduces _acc))
- : M (PTW_Result (mword 56) SV48_PTE))
+ ext_ptw (Z.sub _reclimit 1) _ (_limit_reduces _acc))
+ : M (PTW_Result (mword 56) SV48_PTE)
+ else returnm (PTW_Failure (PTW_Invalid_PTE tt, ext_ptw)))
: M (PTW_Result (mword 56) SV48_PTE)
else
- (checkPTEPermission ac priv mxr do_sum pattr) >>= fun w__3 : bool =>
- returnm ((if sumbool_of_bool ((negb w__3)) then PTW_Failure (PTW_No_Permission)
- else if sumbool_of_bool ((Z.gtb level 0)) then
- let mask :=
- sub_vec_int
- (shiftl
- (xor_vec (_get_SV48_PTE_PPNi pte)
- (xor_vec (_get_SV48_PTE_PPNi pte)
- (EXTZ 44 (vec_of_bits [B1] : mword 1))))
- (Z.mul level SV48_LEVEL_BITS)) 1 in
- if ((neq_vec (and_vec (_get_SV48_PTE_PPNi pte) mask)
- (EXTZ 44 (vec_of_bits [B0] : mword 1)))) then
- PTW_Failure
- (PTW_Misaligned)
+ (checkPTEPermission ac priv mxr do_sum pattr ext_pte ext_ptw) >>= fun w__3 : PTE_Check =>
+ returnm (match w__3 with
+ | PTE_Check_Failure ext_ptw => PTW_Failure (PTW_No_Permission tt, ext_ptw)
+ | PTE_Check_Success ext_ptw =>
+ if sumbool_of_bool (Z.gtb level 0) then
+ let mask :=
+ sub_vec_int
+ (shiftl
+ (xor_vec (_get_SV48_PTE_PPNi pte)
+ (xor_vec (_get_SV48_PTE_PPNi pte) (EXTZ 44 ('b"1" : mword 1))))
+ (Z.mul level SV48_LEVEL_BITS)) 1 in
+ if neq_vec (and_vec (_get_SV48_PTE_PPNi pte) mask)
+ (EXTZ 44 ('b"0" : mword 1)) then
+ PTW_Failure (PTW_Misaligned tt, ext_ptw)
+ else
+ let ppn :=
+ or_vec (_get_SV48_PTE_PPNi pte)
+ (and_vec (EXTZ 44 (_get_SV48_Vaddr_VPNi va)) mask) in
+ PTW_Success
+ (concat_vec ppn (_get_SV48_Vaddr_PgOfs va), pte, pte_addr, build_ex
+ level, is_global, ext_ptw)
else
- let ppn :=
- or_vec (_get_SV48_PTE_PPNi pte)
- (and_vec (EXTZ 44 (_get_SV48_Vaddr_VPNi va)) mask) in
PTW_Success
- ((concat_vec ppn (_get_SV48_Vaddr_PgOfs va), pte, pte_addr, build_ex
- level, is_global))
- else
- PTW_Success
- ((concat_vec (_get_SV48_PTE_PPNi pte) (_get_SV48_Vaddr_PgOfs va), pte, pte_addr, build_ex
- level, is_global)))
- : PTW_Result (mword 56) SV48_PTE))
+ (concat_vec (_get_SV48_PTE_PPNi pte) (_get_SV48_Vaddr_PgOfs va), pte, pte_addr, build_ex
+ level, is_global, ext_ptw)
+ end))
: M (PTW_Result (mword 56) SV48_PTE)
end)
- : M (PTW_Result (mword 56) SV48_PTE).
+ : M (PTW_Result (mword 56) SV48_PTE)
+).
+Defined.
+
Definition walk48
-(_arg0 : mword 48) (_arg1 : AccessType) (_arg2 : Privilege) (_arg3 : bool) (_arg4 : bool)
-(_arg5 : mword 56) (level : Z) (_arg7 : bool) `{ArithFact (0 <= level)}
+(_arg0 : mword 48) (_arg1 : AccessType unit) (_arg2 : Privilege) (_arg3 : bool) (_arg4 : bool)
+(_arg5 : mword 56) (level : Z) (_arg7 : bool) (_arg8 : unit) `{ArithFact (0 <=? level)}
: M (PTW_Result (mword 56) SV48_PTE) :=
-
- (_rec_walk48 _arg0 _arg1 _arg2 _arg3 _arg4 _arg5 level _arg7 (level : Z) (Zwf_guarded _))
+ (_rec_walk48 _arg0 _arg1 _arg2 _arg3 _arg4 _arg5 level _arg7 _arg8 (level : Z) (Zwf_guarded _))
: M (PTW_Result (mword 56) SV48_PTE).
-Definition lookup_TLB48 (asid : mword 16) (vaddr : mword 48)
-: M (option (({n : Z & ArithFact (n >= 0)} * TLB_Entry 16 48 56 64))) :=
-
+Definition lookup_TLB48 (asid : mword 16) (vaddr : mword 48)
+: M (option (({n : Z & ArithFact (n >=? 0)} * TLB_Entry 16 48 56 64))) :=
read_reg tlb48_ref >>= fun w__0 : option (TLB_Entry 16 48 56 64) =>
- returnm ((match w__0 with
- | None => None
- | Some (e) => if ((match_TLB_Entry e asid vaddr)) then Some ((build_ex 0, e)) else None
- end)
- : option (({n : Z & ArithFact (n >= 0)} * TLB_Entry 16 48 56 64))).
+ returnm (match w__0 with
+ | None => None
+ | Some e => if match_TLB_Entry e asid vaddr then Some (build_ex 0, e) else None
+ end).
Definition add_to_TLB48
(asid : mword 16) (vAddr : mword 48) (pAddr : mword 56) (pte : SV48_PTE) (pteAddr : mword 56)
-(level : Z) (global : bool) `{ArithFact (0 <= level)}
+(level : Z) (global : bool) `{ArithFact (0 <=? level)}
: M (unit) :=
-
(make_TLB_Entry asid global vAddr pAddr (_get_SV48_PTE_bits pte) level pteAddr SV48_LEVEL_BITS) >>= fun ent : TLB48_Entry =>
- write_reg tlb48_ref (Some (ent))
+ write_reg tlb48_ref (Some ent)
: M (unit).
-Definition write_TLB48 (idx : Z) (ent : TLB_Entry 16 48 56 64) `{ArithFact (0 <= idx)}
-: M (unit) :=
-
- write_reg tlb48_ref (Some (ent))
- : M (unit).
+Definition write_TLB48 (idx : Z) (ent : TLB_Entry 16 48 56 64) `{ArithFact (0 <=? idx)} : M (unit) :=
+ write_reg tlb48_ref (Some ent) : M (unit).
-Definition flush_TLB48 (asid : option (mword 16)) (addr : option (mword 48))
-: M (unit) :=
-
+Definition flush_TLB48 (asid : option (mword 16)) (addr : option (mword 48)) : M (unit) :=
read_reg tlb48_ref >>= fun w__0 : option (TLB_Entry 16 48 56 64) =>
(match w__0 with
- | None => returnm (tt : unit)
- | Some (e) =>
- (if ((flush_TLB_Entry e asid addr)) then write_reg tlb48_ref None : M (unit)
- else returnm (tt : unit))
+ | None => returnm tt
+ | Some e =>
+ (if flush_TLB_Entry e asid addr then write_reg tlb48_ref None : M (unit)
+ else returnm tt)
: M (unit)
end)
: M (unit).
Definition translate48
-(asid : mword 16) (ptb : mword 56) (vAddr : mword 48) (ac : AccessType) (priv : Privilege)
-(mxr : bool) (do_sum : bool) (level : Z) `{ArithFact (0 <= level)}
+(asid : mword 16) (ptb : mword 56) (vAddr : mword 48) (ac : AccessType unit) (priv : Privilege)
+(mxr : bool) (do_sum : bool) (level : Z) (ext_ptw : unit) `{ArithFact (0 <=? level)}
: M (TR_Result (mword 56) PTW_Error) :=
-
- (walk48 vAddr ac priv mxr do_sum ptb level false) >>= fun w__0 : PTW_Result (mword 56) SV48_PTE =>
+ (walk48 vAddr ac priv mxr do_sum ptb level false ext_ptw) >>= fun w__0 : PTW_Result (mword 56) SV48_PTE =>
(match w__0 with
- | PTW_Failure (f) => returnm ((TR_Failure (f)) : TR_Result (mword 56) PTW_Error)
- | PTW_Success ((pAddr, pte, pteAddr, existT _ level _, global)) =>
- (match (update_PTE_Bits (Mk_PTE_Bits (_get_SV48_PTE_BITS pte)) ac) with
+ | PTW_Failure (f, ext_ptw) => returnm (TR_Failure (f, ext_ptw))
+ | PTW_Success (pAddr, pte, pteAddr, existT _ level _, global, ext_ptw) =>
+ (match (update_PTE_Bits (Mk_PTE_Bits (_get_SV48_PTE_BITS pte)) ac (_get_SV48_PTE_Ext pte)) with
| None =>
(add_to_TLB48 asid vAddr pAddr pte pteAddr level global) >>
- returnm ((TR_Address
- (pAddr))
- : TR_Result (mword 56) PTW_Error)
- | Some (pbits) =>
- (if ((negb (plat_enable_dirty_update tt))) then
- returnm ((TR_Failure
- (PTW_PTE_Update))
- : TR_Result (mword 56) PTW_Error)
+ returnm (TR_Address (pAddr, ext_ptw))
+ | Some (pbits, ext) =>
+ (if negb (plat_enable_dirty_update tt) then
+ returnm (TR_Failure (PTW_PTE_Update tt, ext_ptw))
else
let w_pte : SV48_PTE := _update_SV48_PTE_BITS pte (_get_PTE_Bits_bits pbits) in
+ let w_pte : SV48_PTE := (_update_SV48_PTE_Ext w_pte ext) : SV48_PTE in
(mem_write_value (EXTZ 64 pteAddr) 8 (_get_SV48_PTE_bits w_pte) false false false) >>= fun w__1 : MemoryOpResult bool =>
(match w__1 with
- | MemValue (_) =>
+ | MemValue _ =>
(add_to_TLB48 asid vAddr pAddr w_pte pteAddr level global) >>
- returnm ((TR_Address
- (pAddr))
- : TR_Result (mword 56) PTW_Error)
- | MemException (e) =>
- returnm ((TR_Failure (PTW_Access)) : TR_Result (mword 56) PTW_Error)
+ returnm (TR_Address (pAddr, ext_ptw))
+ | MemException e => returnm (TR_Failure (PTW_Access tt, ext_ptw))
end)
: M (TR_Result (mword 56) PTW_Error))
: M (TR_Result (mword 56) PTW_Error)
@@ -13707,238 +13242,199 @@ Definition translate48
end)
: M (TR_Result (mword 56) PTW_Error).
-Definition init_vmem_sv48 '(tt : unit) : M (unit) := write_reg tlb48_ref None : M (unit).
+Definition init_vmem_sv48 '(tt : unit) : M (unit) := write_reg tlb48_ref None : M (unit).
-Definition legalize_satp (a : Architecture) (o : mword 64) (v : mword 64)
-: mword 64 :=
-
+Definition legalize_satp (a : Architecture) (o : mword 64) (v : mword 64) : mword 64 :=
legalize_satp64 a o v.
-Definition translationMode (priv : Privilege)
-: M (SATPMode) :=
-
- (if ((eq_vec (privLevel_to_bits priv) ((privLevel_to_bits Machine) : mword 2))) then
- returnm (Sbare
- : SATPMode)
+Definition isValidSv39Addr (vAddr : mword 64) : bool :=
+ eq_vec (subrange_vec_dec vAddr 63 39)
+ (if eq_bit (access_vec_dec vAddr 38) B1 then ones (Z.add (Z.sub 63 39) 1)
+ else zeros_implicit (Z.add (Z.sub 63 39) 1)).
+
+Definition isValidSv48Addr (vAddr : mword 64) : bool :=
+ eq_vec (subrange_vec_dec vAddr 63 48)
+ (if eq_bit (access_vec_dec vAddr 47) B1 then ones (Z.add (Z.sub 63 48) 1)
+ else zeros_implicit (Z.add (Z.sub 63 48) 1)).
+
+Definition translationMode (priv : Privilege) : M (SATPMode) :=
+ (if generic_eq priv Machine then returnm Sbare
else
read_reg mstatus_ref >>= fun w__0 : Mstatus =>
let arch := architecture (get_mstatus_SXL w__0) in
(match arch with
- | Some (RV64) =>
+ | Some RV64 =>
((read_reg satp_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
let mbits : satp_mode := _get_Satp64_Mode (Mk_Satp64 w__1) in
(match (satp64Mode_of_bits RV64 mbits) with
- | Some (m) => returnm (m : SATPMode)
+ | Some m => returnm m
| None => (internal_error "invalid RV64 translation mode in satp") : M (SATPMode)
end)
: M (SATPMode)
- | Some (RV32) =>
+ | Some RV32 =>
((read_reg satp_ref) : M (mword 64)) >>= fun w__4 : mword 64 =>
let s := Mk_Satp32 (subrange_vec_dec w__4 31 0) in
- returnm ((if ((eq_vec (_get_Satp32_Mode s) ((bool_to_bits false) : mword 1))) then Sbare
- else Sv32)
- : SATPMode)
+ returnm (if eq_vec (_get_Satp32_Mode s) ('b"0" : mword 1) then Sbare else Sv32)
| _ => (internal_error "unsupported address translation arch") : M (SATPMode)
end)
: M (SATPMode))
: M (SATPMode).
-Definition translateAddr (vAddr : mword 64) (ac : AccessType)
+Definition translateAddr (vAddr : mword 64) (ac : AccessType unit)
: M (TR_Result (mword 64) ExceptionType) :=
-
- (match ac with
- | Execute => read_reg cur_privilege_ref : M (Privilege)
- | _ =>
- read_reg mstatus_ref >>= fun w__1 : Mstatus =>
- read_reg cur_privilege_ref >>= fun w__2 : Privilege =>
- (effectivePrivilege w__1 w__2)
- : M (Privilege)
- end) >>= fun effPriv : Privilege =>
- read_reg mstatus_ref >>= fun w__4 : Mstatus =>
- let mxr : bool := eq_vec (_get_Mstatus_MXR w__4) ((bool_to_bits true) : mword 1) in
- read_reg mstatus_ref >>= fun w__5 : Mstatus =>
- let do_sum : bool := eq_vec (_get_Mstatus_SUM w__5) ((bool_to_bits true) : mword 1) in
+ read_reg mstatus_ref >>= fun w__0 : Mstatus =>
+ read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
+ (effectivePrivilege ac w__0 w__1) >>= fun effPriv : Privilege =>
+ read_reg mstatus_ref >>= fun w__2 : Mstatus =>
+ let mxr : bool := eq_vec (_get_Mstatus_MXR w__2) ('b"1" : mword 1) in
+ read_reg mstatus_ref >>= fun w__3 : Mstatus =>
+ let do_sum : bool := eq_vec (_get_Mstatus_SUM w__3) ('b"1" : mword 1) in
(translationMode effPriv) >>= fun mode : SATPMode =>
- ((read_reg satp_ref) : M (mword 64)) >>= fun w__6 : mword 64 =>
- let asid := curAsid64 w__6 in
- ((read_reg satp_ref) : M (mword 64)) >>= fun w__7 : mword 64 =>
- let ptb := curPTB64 w__7 in
+ ((read_reg satp_ref) : M (mword 64)) >>= fun w__4 : mword 64 =>
+ let asid := curAsid64 w__4 in
+ ((read_reg satp_ref) : M (mword 64)) >>= fun w__5 : mword 64 =>
+ let ptb := curPTB64 w__5 in
+ let ext_ptw : ext_ptw := init_ext_ptw in
(match mode with
- | Sbare => returnm ((TR_Address (vAddr)) : TR_Result (mword 64) ExceptionType)
+ | Sbare => returnm (TR_Address (vAddr, ext_ptw))
| Sv39 =>
- (translate39 asid ptb (subrange_vec_dec vAddr 38 0) ac effPriv mxr do_sum
- (Z.sub SV39_LEVELS 1)) >>= fun w__8 : TR_Result (mword 56) PTW_Error =>
- returnm ((match w__8 with
- | TR_Address (pa) => TR_Address (EXTZ 64 pa)
- | TR_Failure (f) => TR_Failure (translationException ac f)
- end)
- : TR_Result (mword 64) ExceptionType)
+ (if isValidSv39Addr vAddr then
+ (translate39 asid ptb (subrange_vec_dec vAddr 38 0) ac effPriv mxr do_sum
+ (Z.sub SV39_LEVELS 1) ext_ptw) >>= fun w__6 : TR_Result (mword 56) PTW_Error =>
+ returnm (match w__6 with
+ | TR_Address (pa, ext_ptw) => TR_Address (EXTZ 64 pa, ext_ptw)
+ | TR_Failure (f, ext_ptw) => TR_Failure (translationException ac f, ext_ptw)
+ end)
+ else returnm (TR_Failure (translationException ac (PTW_Invalid_Addr tt), ext_ptw)))
+ : M (TR_Result (mword 64) ExceptionType)
| Sv48 =>
- (translate48 asid ptb (subrange_vec_dec vAddr 47 0) ac effPriv mxr do_sum
- (Z.sub SV48_LEVELS 1)) >>= fun w__9 : TR_Result (mword 56) PTW_Error =>
- returnm ((match w__9 with
- | TR_Address (pa) => TR_Address (EXTZ 64 pa)
- | TR_Failure (f) => TR_Failure (translationException ac f)
- end)
- : TR_Result (mword 64) ExceptionType)
+ (if isValidSv48Addr vAddr then
+ (translate48 asid ptb (subrange_vec_dec vAddr 47 0) ac effPriv mxr do_sum
+ (Z.sub SV48_LEVELS 1) ext_ptw) >>= fun w__8 : TR_Result (mword 56) PTW_Error =>
+ returnm (match w__8 with
+ | TR_Address (pa, ext_ptw) => TR_Address (EXTZ 64 pa, ext_ptw)
+ | TR_Failure (f, ext_ptw) => TR_Failure (translationException ac f, ext_ptw)
+ end)
+ else returnm (TR_Failure (translationException ac (PTW_Invalid_Addr tt), ext_ptw)))
+ : M (TR_Result (mword 64) ExceptionType)
| _ =>
(internal_error "unsupported address translation scheme")
: M (TR_Result (mword 64) ExceptionType)
end)
: M (TR_Result (mword 64) ExceptionType).
-Definition flush_TLB (asid_xlen : option (mword 64)) (addr_xlen : option (mword 64))
-: M (unit) :=
-
+Definition flush_TLB (asid_xlen : option (mword 64)) (addr_xlen : option (mword 64)) : M (unit) :=
let '(addr39, addr48) :=
(match addr_xlen with
| None => (None, None)
- | Some (a) => (Some (subrange_vec_dec a 38 0), Some (subrange_vec_dec a 47 0))
+ | Some a => (Some (subrange_vec_dec a 38 0), Some (subrange_vec_dec a 47 0))
end)
: (option vaddr39 * option vaddr48) in
let asid : option asid64 :=
- match asid_xlen with | None => None | Some (a) => Some (subrange_vec_dec a 15 0) end in
+ match asid_xlen with | None => None | Some a => Some (subrange_vec_dec a 15 0) end in
(flush_TLB39 asid addr39) >> (flush_TLB48 asid addr48) : M (unit).
-Definition init_vmem '(tt : unit)
-: M (unit) :=
-
+Definition init_vmem '(tt : unit) : M (unit) :=
(init_vmem_sv39 tt) >> (init_vmem_sv48 tt) : M (unit).
-Definition encdec_uop_forwards (arg_ : uop)
-: mword 7 :=
-
- match arg_ with
- | RISCV_LUI => (vec_of_bits [B0;B1;B1;B0;B1;B1;B1] : mword 7)
- | RISCV_AUIPC => (vec_of_bits [B0;B0;B1;B0;B1;B1;B1] : mword 7)
- end.
+Definition encdec_uop_forwards (arg_ : uop) : mword 7 :=
+ match arg_ with | RISCV_LUI => 'b"0110111" : mword 7 | RISCV_AUIPC => 'b"0010111" : mword 7 end.
-Definition encdec_uop_backwards (arg_ : mword 7)
-: M (uop) :=
-
+Definition encdec_uop_backwards (arg_ : mword 7) : M (uop) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B1;B1;B1] : mword 7))) then
- returnm (RISCV_LUI
- : uop)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B1;B1;B1] : mword 7))) then
- returnm (RISCV_AUIPC
- : uop)
+ (if eq_vec b__0 ('b"0110111" : mword 7) then returnm RISCV_LUI
+ else if eq_vec b__0 ('b"0010111" : mword 7) then returnm RISCV_AUIPC
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (uop).
-Definition encdec_uop_forwards_matches (arg_ : uop)
-: bool :=
-
+Definition encdec_uop_forwards_matches (arg_ : uop) : bool :=
match arg_ with | RISCV_LUI => true | RISCV_AUIPC => true end.
-Definition encdec_uop_backwards_matches (arg_ : mword 7)
-: bool :=
-
+Definition encdec_uop_backwards_matches (arg_ : mword 7) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B1;B1;B1] : mword 7))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B1;B1;B1] : mword 7))) then true
+ if eq_vec b__0 ('b"0110111" : mword 7) then true
+ else if eq_vec b__0 ('b"0010111" : mword 7) then true
else false.
-Definition utype_mnemonic_forwards (arg_ : uop)
-: string :=
-
+Definition utype_mnemonic_forwards (arg_ : uop) : string :=
match arg_ with | RISCV_LUI => "lui" | RISCV_AUIPC => "auipc" end.
-Definition utype_mnemonic_backwards (arg_ : string)
-: M (uop) :=
-
+Definition utype_mnemonic_backwards (arg_ : string) : M (uop) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "lui")) then returnm (RISCV_LUI : uop)
- else if ((generic_eq p0_ "auipc")) then returnm (RISCV_AUIPC : uop)
+ (if generic_eq p0_ "lui" then returnm RISCV_LUI
+ else if generic_eq p0_ "auipc" then returnm RISCV_AUIPC
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (uop).
-Definition utype_mnemonic_forwards_matches (arg_ : uop)
-: bool :=
-
+Definition utype_mnemonic_forwards_matches (arg_ : uop) : bool :=
match arg_ with | RISCV_LUI => true | RISCV_AUIPC => true end.
-Definition utype_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition utype_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "lui")) then true
- else if ((generic_eq p0_ "auipc")) then true
+ if generic_eq p0_ "lui" then true
+ else if generic_eq p0_ "auipc" then true
else false.
-Definition _s496_ (_s497_ : string)
-: option string :=
-
- let _s498_ := _s497_ in
- if ((string_startswith _s498_ "auipc")) then
- match (string_drop _s498_ (projT1 (string_length "auipc"))) with | s_ => Some (s_) end
+Definition _s757_ (_s758_ : string) : option string :=
+ let _s759_ := _s758_ in
+ if string_startswith _s759_ "auipc" then
+ match (string_drop _s759_ (projT1 (string_length "auipc"))) with | s_ => Some s_ end
else None.
-Definition _s492_ (_s493_ : string)
-: option string :=
-
- let _s494_ := _s493_ in
- if ((string_startswith _s494_ "lui")) then
- match (string_drop _s494_ (projT1 (string_length "lui"))) with | s_ => Some (s_) end
- else None.
-
-Definition utype_mnemonic_matches_prefix (arg_ : string)
-: M (option ((uop * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s495_ := arg_ in
- (if ((match (_s492_ _s495_) with | Some (s_) => true | _ => false end)) then
- (match (_s492_ _s495_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_LUI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((uop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((uop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((uop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s496_ _s495_) with | Some (s_) => true | _ => false end)) then
- (match (_s496_ _s495_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_AUIPC, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((uop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((uop * {n : Z & ArithFact (n >= 0)})))
+Definition _s753_ (_s754_ : string) : option string :=
+ let _s755_ := _s754_ in
+ if string_startswith _s755_ "lui" then
+ match (string_drop _s755_ (projT1 (string_length "lui"))) with | s_ => Some s_ end
+ else None.
+
+Definition utype_mnemonic_matches_prefix (arg_ : string)
+: M (option ((uop * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s756_ := arg_ in
+ (if match (_s753_ _s756_) with | Some s_ => true | _ => false end then
+ (match (_s753_ _s756_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_LUI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((uop * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((uop * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((uop * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((uop * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((uop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s757_ _s756_) with | Some s_ => true | _ => false end then
+ (match (_s757_ _s756_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_AUIPC, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((uop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((uop * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((uop * {n : Z & ArithFact (n >=? 0)}))).
-Definition encdec_bop_forwards (arg_ : bop)
-: mword 3 :=
-
+Definition encdec_bop_forwards (arg_ : bop) : mword 3 :=
match arg_ with
- | RISCV_BEQ => (vec_of_bits [B0;B0;B0] : mword 3)
- | RISCV_BNE => (vec_of_bits [B0;B0;B1] : mword 3)
- | RISCV_BLT => (vec_of_bits [B1;B0;B0] : mword 3)
- | RISCV_BGE => (vec_of_bits [B1;B0;B1] : mword 3)
- | RISCV_BLTU => (vec_of_bits [B1;B1;B0] : mword 3)
- | RISCV_BGEU => (vec_of_bits [B1;B1;B1] : mword 3)
+ | RISCV_BEQ => 'b"000" : mword 3
+ | RISCV_BNE => 'b"001" : mword 3
+ | RISCV_BLT => 'b"100" : mword 3
+ | RISCV_BGE => 'b"101" : mword 3
+ | RISCV_BLTU => 'b"110" : mword 3
+ | RISCV_BGEU => 'b"111" : mword 3
end.
-Definition encdec_bop_backwards (arg_ : mword 3)
-: M (bop) :=
-
+Definition encdec_bop_backwards (arg_ : mword 3) : M (bop) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then returnm (RISCV_BEQ : bop)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then returnm (RISCV_BNE : bop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then returnm (RISCV_BLT : bop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then returnm (RISCV_BGE : bop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then returnm (RISCV_BLTU : bop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1] : mword 3))) then returnm (RISCV_BGEU : bop)
+ (if eq_vec b__0 ('b"000" : mword 3) then returnm RISCV_BEQ
+ else if eq_vec b__0 ('b"001" : mword 3) then returnm RISCV_BNE
+ else if eq_vec b__0 ('b"100" : mword 3) then returnm RISCV_BLT
+ else if eq_vec b__0 ('b"101" : mword 3) then returnm RISCV_BGE
+ else if eq_vec b__0 ('b"110" : mword 3) then returnm RISCV_BLTU
+ else if eq_vec b__0 ('b"111" : mword 3) then returnm RISCV_BGEU
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bop).
-Definition encdec_bop_forwards_matches (arg_ : bop)
-: bool :=
-
+Definition encdec_bop_forwards_matches (arg_ : bop) : bool :=
match arg_ with
| RISCV_BEQ => true
| RISCV_BNE => true
@@ -13948,21 +13444,17 @@ Definition encdec_bop_forwards_matches (arg_ : bop)
| RISCV_BGEU => true
end.
-Definition encdec_bop_backwards_matches (arg_ : mword 3)
-: bool :=
-
+Definition encdec_bop_backwards_matches (arg_ : mword 3) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1] : mword 3))) then true
+ if eq_vec b__0 ('b"000" : mword 3) then true
+ else if eq_vec b__0 ('b"001" : mword 3) then true
+ else if eq_vec b__0 ('b"100" : mword 3) then true
+ else if eq_vec b__0 ('b"101" : mword 3) then true
+ else if eq_vec b__0 ('b"110" : mword 3) then true
+ else if eq_vec b__0 ('b"111" : mword 3) then true
else false.
-Definition btype_mnemonic_forwards (arg_ : bop)
-: string :=
-
+Definition btype_mnemonic_forwards (arg_ : bop) : string :=
match arg_ with
| RISCV_BEQ => "beq"
| RISCV_BNE => "bne"
@@ -13972,22 +13464,18 @@ Definition btype_mnemonic_forwards (arg_ : bop)
| RISCV_BGEU => "bgeu"
end.
-Definition btype_mnemonic_backwards (arg_ : string)
-: M (bop) :=
-
+Definition btype_mnemonic_backwards (arg_ : string) : M (bop) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "beq")) then returnm (RISCV_BEQ : bop)
- else if ((generic_eq p0_ "bne")) then returnm (RISCV_BNE : bop)
- else if ((generic_eq p0_ "blt")) then returnm (RISCV_BLT : bop)
- else if ((generic_eq p0_ "bge")) then returnm (RISCV_BGE : bop)
- else if ((generic_eq p0_ "bltu")) then returnm (RISCV_BLTU : bop)
- else if ((generic_eq p0_ "bgeu")) then returnm (RISCV_BGEU : bop)
+ (if generic_eq p0_ "beq" then returnm RISCV_BEQ
+ else if generic_eq p0_ "bne" then returnm RISCV_BNE
+ else if generic_eq p0_ "blt" then returnm RISCV_BLT
+ else if generic_eq p0_ "bge" then returnm RISCV_BGE
+ else if generic_eq p0_ "bltu" then returnm RISCV_BLTU
+ else if generic_eq p0_ "bgeu" then returnm RISCV_BGEU
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bop).
-Definition btype_mnemonic_forwards_matches (arg_ : bop)
-: bool :=
-
+Definition btype_mnemonic_forwards_matches (arg_ : bop) : bool :=
match arg_ with
| RISCV_BEQ => true
| RISCV_BNE => true
@@ -13997,167 +13485,140 @@ Definition btype_mnemonic_forwards_matches (arg_ : bop)
| RISCV_BGEU => true
end.
-Definition btype_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition btype_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "beq")) then true
- else if ((generic_eq p0_ "bne")) then true
- else if ((generic_eq p0_ "blt")) then true
- else if ((generic_eq p0_ "bge")) then true
- else if ((generic_eq p0_ "bltu")) then true
- else if ((generic_eq p0_ "bgeu")) then true
+ if generic_eq p0_ "beq" then true
+ else if generic_eq p0_ "bne" then true
+ else if generic_eq p0_ "blt" then true
+ else if generic_eq p0_ "bge" then true
+ else if generic_eq p0_ "bltu" then true
+ else if generic_eq p0_ "bgeu" then true
else false.
-Definition _s520_ (_s521_ : string)
-: option string :=
-
- let _s522_ := _s521_ in
- if ((string_startswith _s522_ "bgeu")) then
- match (string_drop _s522_ (projT1 (string_length "bgeu"))) with | s_ => Some (s_) end
+Definition _s781_ (_s782_ : string) : option string :=
+ let _s783_ := _s782_ in
+ if string_startswith _s783_ "bgeu" then
+ match (string_drop _s783_ (projT1 (string_length "bgeu"))) with | s_ => Some s_ end
else None.
-Definition _s516_ (_s517_ : string)
-: option string :=
-
- let _s518_ := _s517_ in
- if ((string_startswith _s518_ "bltu")) then
- match (string_drop _s518_ (projT1 (string_length "bltu"))) with | s_ => Some (s_) end
+Definition _s777_ (_s778_ : string) : option string :=
+ let _s779_ := _s778_ in
+ if string_startswith _s779_ "bltu" then
+ match (string_drop _s779_ (projT1 (string_length "bltu"))) with | s_ => Some s_ end
else None.
-Definition _s512_ (_s513_ : string)
-: option string :=
-
- let _s514_ := _s513_ in
- if ((string_startswith _s514_ "bge")) then
- match (string_drop _s514_ (projT1 (string_length "bge"))) with | s_ => Some (s_) end
+Definition _s773_ (_s774_ : string) : option string :=
+ let _s775_ := _s774_ in
+ if string_startswith _s775_ "bge" then
+ match (string_drop _s775_ (projT1 (string_length "bge"))) with | s_ => Some s_ end
else None.
-Definition _s508_ (_s509_ : string)
-: option string :=
-
- let _s510_ := _s509_ in
- if ((string_startswith _s510_ "blt")) then
- match (string_drop _s510_ (projT1 (string_length "blt"))) with | s_ => Some (s_) end
+Definition _s769_ (_s770_ : string) : option string :=
+ let _s771_ := _s770_ in
+ if string_startswith _s771_ "blt" then
+ match (string_drop _s771_ (projT1 (string_length "blt"))) with | s_ => Some s_ end
else None.
-Definition _s504_ (_s505_ : string)
-: option string :=
-
- let _s506_ := _s505_ in
- if ((string_startswith _s506_ "bne")) then
- match (string_drop _s506_ (projT1 (string_length "bne"))) with | s_ => Some (s_) end
+Definition _s765_ (_s766_ : string) : option string :=
+ let _s767_ := _s766_ in
+ if string_startswith _s767_ "bne" then
+ match (string_drop _s767_ (projT1 (string_length "bne"))) with | s_ => Some s_ end
else None.
-Definition _s500_ (_s501_ : string)
-: option string :=
-
- let _s502_ := _s501_ in
- if ((string_startswith _s502_ "beq")) then
- match (string_drop _s502_ (projT1 (string_length "beq"))) with | s_ => Some (s_) end
- else None.
-
-Definition btype_mnemonic_matches_prefix (arg_ : string)
-: M (option ((bop * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s503_ := arg_ in
- (if ((match (_s500_ _s503_) with | Some (s_) => true | _ => false end)) then
- (match (_s500_ _s503_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_BEQ, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s504_ _s503_) with | Some (s_) => true | _ => false end)) then
- (match (_s504_ _s503_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_BNE, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s508_ _s503_) with | Some (s_) => true | _ => false end)) then
- (match (_s508_ _s503_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_BLT, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s512_ _s503_) with | Some (s_) => true | _ => false end)) then
- (match (_s512_ _s503_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_BGE, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s516_ _s503_) with | Some (s_) => true | _ => false end)) then
- (match (_s516_ _s503_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_BLTU, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s520_ _s503_) with | Some (s_) => true | _ => false end)) then
- (match (_s520_ _s503_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_BGEU, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
+Definition _s761_ (_s762_ : string) : option string :=
+ let _s763_ := _s762_ in
+ if string_startswith _s763_ "beq" then
+ match (string_drop _s763_ (projT1 (string_length "beq"))) with | s_ => Some s_ end
+ else None.
+
+Definition btype_mnemonic_matches_prefix (arg_ : string)
+: M (option ((bop * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s764_ := arg_ in
+ (if match (_s761_ _s764_) with | Some s_ => true | _ => false end then
+ (match (_s761_ _s764_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_BEQ, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s765_ _s764_) with | Some s_ => true | _ => false end then
+ (match (_s765_ _s764_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_BNE, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((bop * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((bop * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((bop * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s769_ _s764_) with | Some s_ => true | _ => false end then
+ (match (_s769_ _s764_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_BLT, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s773_ _s764_) with | Some s_ => true | _ => false end then
+ (match (_s773_ _s764_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_BGE, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s777_ _s764_) with | Some s_ => true | _ => false end then
+ (match (_s777_ _s764_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_BLTU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s781_ _s764_) with | Some s_ => true | _ => false end then
+ (match (_s781_ _s764_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_BGEU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bop * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((bop * {n : Z & ArithFact (n >=? 0)}))).
-Definition encdec_iop_forwards (arg_ : iop)
-: mword 3 :=
-
+Definition encdec_iop_forwards (arg_ : iop) : mword 3 :=
match arg_ with
- | RISCV_ADDI => (vec_of_bits [B0;B0;B0] : mword 3)
- | RISCV_SLTI => (vec_of_bits [B0;B1;B0] : mword 3)
- | RISCV_SLTIU => (vec_of_bits [B0;B1;B1] : mword 3)
- | RISCV_ANDI => (vec_of_bits [B1;B1;B1] : mword 3)
- | RISCV_ORI => (vec_of_bits [B1;B1;B0] : mword 3)
- | RISCV_XORI => (vec_of_bits [B1;B0;B0] : mword 3)
+ | RISCV_ADDI => 'b"000" : mword 3
+ | RISCV_SLTI => 'b"010" : mword 3
+ | RISCV_SLTIU => 'b"011" : mword 3
+ | RISCV_ANDI => 'b"111" : mword 3
+ | RISCV_ORI => 'b"110" : mword 3
+ | RISCV_XORI => 'b"100" : mword 3
end.
-Definition encdec_iop_backwards (arg_ : mword 3)
-: M (iop) :=
-
+Definition encdec_iop_backwards (arg_ : mword 3) : M (iop) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then returnm (RISCV_ADDI : iop)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then returnm (RISCV_SLTI : iop)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then returnm (RISCV_SLTIU : iop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1] : mword 3))) then returnm (RISCV_ANDI : iop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then returnm (RISCV_ORI : iop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then returnm (RISCV_XORI : iop)
+ (if eq_vec b__0 ('b"000" : mword 3) then returnm RISCV_ADDI
+ else if eq_vec b__0 ('b"010" : mword 3) then returnm RISCV_SLTI
+ else if eq_vec b__0 ('b"011" : mword 3) then returnm RISCV_SLTIU
+ else if eq_vec b__0 ('b"111" : mword 3) then returnm RISCV_ANDI
+ else if eq_vec b__0 ('b"110" : mword 3) then returnm RISCV_ORI
+ else if eq_vec b__0 ('b"100" : mword 3) then returnm RISCV_XORI
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (iop).
-Definition encdec_iop_forwards_matches (arg_ : iop)
-: bool :=
-
+Definition encdec_iop_forwards_matches (arg_ : iop) : bool :=
match arg_ with
| RISCV_ADDI => true
| RISCV_SLTI => true
@@ -14167,21 +13628,17 @@ Definition encdec_iop_forwards_matches (arg_ : iop)
| RISCV_XORI => true
end.
-Definition encdec_iop_backwards_matches (arg_ : mword 3)
-: bool :=
-
+Definition encdec_iop_backwards_matches (arg_ : mword 3) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0] : mword 3))) then true
+ if eq_vec b__0 ('b"000" : mword 3) then true
+ else if eq_vec b__0 ('b"010" : mword 3) then true
+ else if eq_vec b__0 ('b"011" : mword 3) then true
+ else if eq_vec b__0 ('b"111" : mword 3) then true
+ else if eq_vec b__0 ('b"110" : mword 3) then true
+ else if eq_vec b__0 ('b"100" : mword 3) then true
else false.
-Definition itype_mnemonic_forwards (arg_ : iop)
-: string :=
-
+Definition itype_mnemonic_forwards (arg_ : iop) : string :=
match arg_ with
| RISCV_ADDI => "addi"
| RISCV_SLTI => "slti"
@@ -14191,22 +13648,18 @@ Definition itype_mnemonic_forwards (arg_ : iop)
| RISCV_ANDI => "andi"
end.
-Definition itype_mnemonic_backwards (arg_ : string)
-: M (iop) :=
-
+Definition itype_mnemonic_backwards (arg_ : string) : M (iop) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "addi")) then returnm (RISCV_ADDI : iop)
- else if ((generic_eq p0_ "slti")) then returnm (RISCV_SLTI : iop)
- else if ((generic_eq p0_ "sltiu")) then returnm (RISCV_SLTIU : iop)
- else if ((generic_eq p0_ "xori")) then returnm (RISCV_XORI : iop)
- else if ((generic_eq p0_ "ori")) then returnm (RISCV_ORI : iop)
- else if ((generic_eq p0_ "andi")) then returnm (RISCV_ANDI : iop)
+ (if generic_eq p0_ "addi" then returnm RISCV_ADDI
+ else if generic_eq p0_ "slti" then returnm RISCV_SLTI
+ else if generic_eq p0_ "sltiu" then returnm RISCV_SLTIU
+ else if generic_eq p0_ "xori" then returnm RISCV_XORI
+ else if generic_eq p0_ "ori" then returnm RISCV_ORI
+ else if generic_eq p0_ "andi" then returnm RISCV_ANDI
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (iop).
-Definition itype_mnemonic_forwards_matches (arg_ : iop)
-: bool :=
-
+Definition itype_mnemonic_forwards_matches (arg_ : iop) : bool :=
match arg_ with
| RISCV_ADDI => true
| RISCV_SLTI => true
@@ -14216,268 +13669,219 @@ Definition itype_mnemonic_forwards_matches (arg_ : iop)
| RISCV_ANDI => true
end.
-Definition itype_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition itype_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "addi")) then true
- else if ((generic_eq p0_ "slti")) then true
- else if ((generic_eq p0_ "sltiu")) then true
- else if ((generic_eq p0_ "xori")) then true
- else if ((generic_eq p0_ "ori")) then true
- else if ((generic_eq p0_ "andi")) then true
+ if generic_eq p0_ "addi" then true
+ else if generic_eq p0_ "slti" then true
+ else if generic_eq p0_ "sltiu" then true
+ else if generic_eq p0_ "xori" then true
+ else if generic_eq p0_ "ori" then true
+ else if generic_eq p0_ "andi" then true
else false.
-Definition _s544_ (_s545_ : string)
-: option string :=
-
- let _s546_ := _s545_ in
- if ((string_startswith _s546_ "andi")) then
- match (string_drop _s546_ (projT1 (string_length "andi"))) with | s_ => Some (s_) end
+Definition _s805_ (_s806_ : string) : option string :=
+ let _s807_ := _s806_ in
+ if string_startswith _s807_ "andi" then
+ match (string_drop _s807_ (projT1 (string_length "andi"))) with | s_ => Some s_ end
else None.
-Definition _s540_ (_s541_ : string)
-: option string :=
-
- let _s542_ := _s541_ in
- if ((string_startswith _s542_ "ori")) then
- match (string_drop _s542_ (projT1 (string_length "ori"))) with | s_ => Some (s_) end
+Definition _s801_ (_s802_ : string) : option string :=
+ let _s803_ := _s802_ in
+ if string_startswith _s803_ "ori" then
+ match (string_drop _s803_ (projT1 (string_length "ori"))) with | s_ => Some s_ end
else None.
-Definition _s536_ (_s537_ : string)
-: option string :=
-
- let _s538_ := _s537_ in
- if ((string_startswith _s538_ "xori")) then
- match (string_drop _s538_ (projT1 (string_length "xori"))) with | s_ => Some (s_) end
+Definition _s797_ (_s798_ : string) : option string :=
+ let _s799_ := _s798_ in
+ if string_startswith _s799_ "xori" then
+ match (string_drop _s799_ (projT1 (string_length "xori"))) with | s_ => Some s_ end
else None.
-Definition _s532_ (_s533_ : string)
-: option string :=
-
- let _s534_ := _s533_ in
- if ((string_startswith _s534_ "sltiu")) then
- match (string_drop _s534_ (projT1 (string_length "sltiu"))) with | s_ => Some (s_) end
+Definition _s793_ (_s794_ : string) : option string :=
+ let _s795_ := _s794_ in
+ if string_startswith _s795_ "sltiu" then
+ match (string_drop _s795_ (projT1 (string_length "sltiu"))) with | s_ => Some s_ end
else None.
-Definition _s528_ (_s529_ : string)
-: option string :=
-
- let _s530_ := _s529_ in
- if ((string_startswith _s530_ "slti")) then
- match (string_drop _s530_ (projT1 (string_length "slti"))) with | s_ => Some (s_) end
+Definition _s789_ (_s790_ : string) : option string :=
+ let _s791_ := _s790_ in
+ if string_startswith _s791_ "slti" then
+ match (string_drop _s791_ (projT1 (string_length "slti"))) with | s_ => Some s_ end
else None.
-Definition _s524_ (_s525_ : string)
-: option string :=
-
- let _s526_ := _s525_ in
- if ((string_startswith _s526_ "addi")) then
- match (string_drop _s526_ (projT1 (string_length "addi"))) with | s_ => Some (s_) end
- else None.
-
-Definition itype_mnemonic_matches_prefix (arg_ : string)
-: M (option ((iop * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s527_ := arg_ in
- (if ((match (_s524_ _s527_) with | Some (s_) => true | _ => false end)) then
- (match (_s524_ _s527_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_ADDI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((iop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s528_ _s527_) with | Some (s_) => true | _ => false end)) then
- (match (_s528_ _s527_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLTI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((iop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s532_ _s527_) with | Some (s_) => true | _ => false end)) then
- (match (_s532_ _s527_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLTIU, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((iop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s536_ _s527_) with | Some (s_) => true | _ => false end)) then
- (match (_s536_ _s527_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_XORI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((iop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s540_ _s527_) with | Some (s_) => true | _ => false end)) then
- (match (_s540_ _s527_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_ORI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((iop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s544_ _s527_) with | Some (s_) => true | _ => false end)) then
- (match (_s544_ _s527_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_ANDI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((iop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
+Definition _s785_ (_s786_ : string) : option string :=
+ let _s787_ := _s786_ in
+ if string_startswith _s787_ "addi" then
+ match (string_drop _s787_ (projT1 (string_length "addi"))) with | s_ => Some s_ end
+ else None.
+
+Definition itype_mnemonic_matches_prefix (arg_ : string)
+: M (option ((iop * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s788_ := arg_ in
+ (if match (_s785_ _s788_) with | Some s_ => true | _ => false end then
+ (match (_s785_ _s788_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_ADDI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s789_ _s788_) with | Some s_ => true | _ => false end then
+ (match (_s789_ _s788_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLTI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s793_ _s788_) with | Some s_ => true | _ => false end then
+ (match (_s793_ _s788_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLTIU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s797_ _s788_) with | Some s_ => true | _ => false end then
+ (match (_s797_ _s788_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_XORI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((iop * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((iop * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((iop * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s801_ _s788_) with | Some s_ => true | _ => false end then
+ (match (_s801_ _s788_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_ORI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s805_ _s788_) with | Some s_ => true | _ => false end then
+ (match (_s805_ _s788_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_ANDI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((iop * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((iop * {n : Z & ArithFact (n >=? 0)}))).
-Definition encdec_sop_forwards (arg_ : sop)
-: mword 3 :=
-
+Definition encdec_sop_forwards (arg_ : sop) : mword 3 :=
match arg_ with
- | RISCV_SLLI => (vec_of_bits [B0;B0;B1] : mword 3)
- | RISCV_SRLI => (vec_of_bits [B1;B0;B1] : mword 3)
- | RISCV_SRAI => (vec_of_bits [B1;B0;B1] : mword 3)
+ | RISCV_SLLI => 'b"001" : mword 3
+ | RISCV_SRLI => 'b"101" : mword 3
+ | RISCV_SRAI => 'b"101" : mword 3
end.
-Definition encdec_sop_backwards (arg_ : mword 3)
-: M (sop) :=
-
+Definition encdec_sop_backwards (arg_ : mword 3) : M (sop) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then returnm (RISCV_SLLI : sop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then returnm (RISCV_SRLI : sop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then returnm (RISCV_SRAI : sop)
+ (if eq_vec b__0 ('b"001" : mword 3) then returnm RISCV_SLLI
+ else if eq_vec b__0 ('b"101" : mword 3) then returnm RISCV_SRLI
+ else if eq_vec b__0 ('b"101" : mword 3) then returnm RISCV_SRAI
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (sop).
-Definition encdec_sop_forwards_matches (arg_ : sop)
-: bool :=
-
+Definition encdec_sop_forwards_matches (arg_ : sop) : bool :=
match arg_ with | RISCV_SLLI => true | RISCV_SRLI => true | RISCV_SRAI => true end.
-Definition encdec_sop_backwards_matches (arg_ : mword 3)
-: bool :=
-
+Definition encdec_sop_backwards_matches (arg_ : mword 3) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1] : mword 3))) then true
+ if eq_vec b__0 ('b"001" : mword 3) then true
+ else if eq_vec b__0 ('b"101" : mword 3) then true
+ else if eq_vec b__0 ('b"101" : mword 3) then true
else false.
-Definition shiftiop_mnemonic_forwards (arg_ : sop)
-: string :=
-
+Definition shiftiop_mnemonic_forwards (arg_ : sop) : string :=
match arg_ with | RISCV_SLLI => "slli" | RISCV_SRLI => "srli" | RISCV_SRAI => "srai" end.
-Definition shiftiop_mnemonic_backwards (arg_ : string)
-: M (sop) :=
-
+Definition shiftiop_mnemonic_backwards (arg_ : string) : M (sop) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "slli")) then returnm (RISCV_SLLI : sop)
- else if ((generic_eq p0_ "srli")) then returnm (RISCV_SRLI : sop)
- else if ((generic_eq p0_ "srai")) then returnm (RISCV_SRAI : sop)
+ (if generic_eq p0_ "slli" then returnm RISCV_SLLI
+ else if generic_eq p0_ "srli" then returnm RISCV_SRLI
+ else if generic_eq p0_ "srai" then returnm RISCV_SRAI
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (sop).
-Definition shiftiop_mnemonic_forwards_matches (arg_ : sop)
-: bool :=
-
+Definition shiftiop_mnemonic_forwards_matches (arg_ : sop) : bool :=
match arg_ with | RISCV_SLLI => true | RISCV_SRLI => true | RISCV_SRAI => true end.
-Definition shiftiop_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition shiftiop_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "slli")) then true
- else if ((generic_eq p0_ "srli")) then true
- else if ((generic_eq p0_ "srai")) then true
+ if generic_eq p0_ "slli" then true
+ else if generic_eq p0_ "srli" then true
+ else if generic_eq p0_ "srai" then true
else false.
-Definition _s556_ (_s557_ : string)
-: option string :=
-
- let _s558_ := _s557_ in
- if ((string_startswith _s558_ "srai")) then
- match (string_drop _s558_ (projT1 (string_length "srai"))) with | s_ => Some (s_) end
+Definition _s817_ (_s818_ : string) : option string :=
+ let _s819_ := _s818_ in
+ if string_startswith _s819_ "srai" then
+ match (string_drop _s819_ (projT1 (string_length "srai"))) with | s_ => Some s_ end
else None.
-Definition _s552_ (_s553_ : string)
-: option string :=
-
- let _s554_ := _s553_ in
- if ((string_startswith _s554_ "srli")) then
- match (string_drop _s554_ (projT1 (string_length "srli"))) with | s_ => Some (s_) end
+Definition _s813_ (_s814_ : string) : option string :=
+ let _s815_ := _s814_ in
+ if string_startswith _s815_ "srli" then
+ match (string_drop _s815_ (projT1 (string_length "srli"))) with | s_ => Some s_ end
else None.
-Definition _s548_ (_s549_ : string)
-: option string :=
-
- let _s550_ := _s549_ in
- if ((string_startswith _s550_ "slli")) then
- match (string_drop _s550_ (projT1 (string_length "slli"))) with | s_ => Some (s_) end
- else None.
-
-Definition shiftiop_mnemonic_matches_prefix (arg_ : string)
-: M (option ((sop * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s551_ := arg_ in
- (if ((match (_s548_ _s551_) with | Some (s_) => true | _ => false end)) then
- (match (_s548_ _s551_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLLI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s552_ _s551_) with | Some (s_) => true | _ => false end)) then
- (match (_s552_ _s551_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRLI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s556_ _s551_) with | Some (s_) => true | _ => false end)) then
- (match (_s556_ _s551_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRAI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
+Definition _s809_ (_s810_ : string) : option string :=
+ let _s811_ := _s810_ in
+ if string_startswith _s811_ "slli" then
+ match (string_drop _s811_ (projT1 (string_length "slli"))) with | s_ => Some s_ end
+ else None.
+
+Definition shiftiop_mnemonic_matches_prefix (arg_ : string)
+: M (option ((sop * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s812_ := arg_ in
+ (if match (_s809_ _s812_) with | Some s_ => true | _ => false end then
+ (match (_s809_ _s812_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLLI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s813_ _s812_) with | Some s_ => true | _ => false end then
+ (match (_s813_ _s812_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRLI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s817_ _s812_) with | Some s_ => true | _ => false end then
+ (match (_s817_ _s812_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRAI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((sop * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((sop * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((sop * {n : Z & ArithFact (n >=? 0)}))).
-Definition rtype_mnemonic_forwards (arg_ : rop)
-: string :=
-
+Definition rtype_mnemonic_forwards (arg_ : rop) : string :=
match arg_ with
| RISCV_ADD => "add"
| RISCV_SLT => "slt"
@@ -14491,26 +13895,22 @@ Definition rtype_mnemonic_forwards (arg_ : rop)
| RISCV_SRA => "sra"
end.
-Definition rtype_mnemonic_backwards (arg_ : string)
-: M (rop) :=
-
+Definition rtype_mnemonic_backwards (arg_ : string) : M (rop) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "add")) then returnm (RISCV_ADD : rop)
- else if ((generic_eq p0_ "slt")) then returnm (RISCV_SLT : rop)
- else if ((generic_eq p0_ "sltu")) then returnm (RISCV_SLTU : rop)
- else if ((generic_eq p0_ "and")) then returnm (RISCV_AND : rop)
- else if ((generic_eq p0_ "or")) then returnm (RISCV_OR : rop)
- else if ((generic_eq p0_ "xor")) then returnm (RISCV_XOR : rop)
- else if ((generic_eq p0_ "sll")) then returnm (RISCV_SLL : rop)
- else if ((generic_eq p0_ "srl")) then returnm (RISCV_SRL : rop)
- else if ((generic_eq p0_ "sub")) then returnm (RISCV_SUB : rop)
- else if ((generic_eq p0_ "sra")) then returnm (RISCV_SRA : rop)
+ (if generic_eq p0_ "add" then returnm RISCV_ADD
+ else if generic_eq p0_ "slt" then returnm RISCV_SLT
+ else if generic_eq p0_ "sltu" then returnm RISCV_SLTU
+ else if generic_eq p0_ "and" then returnm RISCV_AND
+ else if generic_eq p0_ "or" then returnm RISCV_OR
+ else if generic_eq p0_ "xor" then returnm RISCV_XOR
+ else if generic_eq p0_ "sll" then returnm RISCV_SLL
+ else if generic_eq p0_ "srl" then returnm RISCV_SRL
+ else if generic_eq p0_ "sub" then returnm RISCV_SUB
+ else if generic_eq p0_ "sra" then returnm RISCV_SRA
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (rop).
-Definition rtype_mnemonic_forwards_matches (arg_ : rop)
-: bool :=
-
+Definition rtype_mnemonic_forwards_matches (arg_ : rop) : bool :=
match arg_ with
| RISCV_ADD => true
| RISCV_SLT => true
@@ -14524,620 +13924,464 @@ Definition rtype_mnemonic_forwards_matches (arg_ : rop)
| RISCV_SRA => true
end.
-Definition rtype_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition rtype_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "add")) then true
- else if ((generic_eq p0_ "slt")) then true
- else if ((generic_eq p0_ "sltu")) then true
- else if ((generic_eq p0_ "and")) then true
- else if ((generic_eq p0_ "or")) then true
- else if ((generic_eq p0_ "xor")) then true
- else if ((generic_eq p0_ "sll")) then true
- else if ((generic_eq p0_ "srl")) then true
- else if ((generic_eq p0_ "sub")) then true
- else if ((generic_eq p0_ "sra")) then true
+ if generic_eq p0_ "add" then true
+ else if generic_eq p0_ "slt" then true
+ else if generic_eq p0_ "sltu" then true
+ else if generic_eq p0_ "and" then true
+ else if generic_eq p0_ "or" then true
+ else if generic_eq p0_ "xor" then true
+ else if generic_eq p0_ "sll" then true
+ else if generic_eq p0_ "srl" then true
+ else if generic_eq p0_ "sub" then true
+ else if generic_eq p0_ "sra" then true
else false.
-Definition _s596_ (_s597_ : string)
-: option string :=
-
- let _s598_ := _s597_ in
- if ((string_startswith _s598_ "sra")) then
- match (string_drop _s598_ (projT1 (string_length "sra"))) with | s_ => Some (s_) end
+Definition _s857_ (_s858_ : string) : option string :=
+ let _s859_ := _s858_ in
+ if string_startswith _s859_ "sra" then
+ match (string_drop _s859_ (projT1 (string_length "sra"))) with | s_ => Some s_ end
else None.
-Definition _s592_ (_s593_ : string)
-: option string :=
-
- let _s594_ := _s593_ in
- if ((string_startswith _s594_ "sub")) then
- match (string_drop _s594_ (projT1 (string_length "sub"))) with | s_ => Some (s_) end
+Definition _s853_ (_s854_ : string) : option string :=
+ let _s855_ := _s854_ in
+ if string_startswith _s855_ "sub" then
+ match (string_drop _s855_ (projT1 (string_length "sub"))) with | s_ => Some s_ end
else None.
-Definition _s588_ (_s589_ : string)
-: option string :=
-
- let _s590_ := _s589_ in
- if ((string_startswith _s590_ "srl")) then
- match (string_drop _s590_ (projT1 (string_length "srl"))) with | s_ => Some (s_) end
+Definition _s849_ (_s850_ : string) : option string :=
+ let _s851_ := _s850_ in
+ if string_startswith _s851_ "srl" then
+ match (string_drop _s851_ (projT1 (string_length "srl"))) with | s_ => Some s_ end
else None.
-Definition _s584_ (_s585_ : string)
-: option string :=
-
- let _s586_ := _s585_ in
- if ((string_startswith _s586_ "sll")) then
- match (string_drop _s586_ (projT1 (string_length "sll"))) with | s_ => Some (s_) end
+Definition _s845_ (_s846_ : string) : option string :=
+ let _s847_ := _s846_ in
+ if string_startswith _s847_ "sll" then
+ match (string_drop _s847_ (projT1 (string_length "sll"))) with | s_ => Some s_ end
else None.
-Definition _s580_ (_s581_ : string)
-: option string :=
-
- let _s582_ := _s581_ in
- if ((string_startswith _s582_ "xor")) then
- match (string_drop _s582_ (projT1 (string_length "xor"))) with | s_ => Some (s_) end
+Definition _s841_ (_s842_ : string) : option string :=
+ let _s843_ := _s842_ in
+ if string_startswith _s843_ "xor" then
+ match (string_drop _s843_ (projT1 (string_length "xor"))) with | s_ => Some s_ end
else None.
-Definition _s576_ (_s577_ : string)
-: option string :=
-
- let _s578_ := _s577_ in
- if ((string_startswith _s578_ "or")) then
- match (string_drop _s578_ (projT1 (string_length "or"))) with | s_ => Some (s_) end
+Definition _s837_ (_s838_ : string) : option string :=
+ let _s839_ := _s838_ in
+ if string_startswith _s839_ "or" then
+ match (string_drop _s839_ (projT1 (string_length "or"))) with | s_ => Some s_ end
else None.
-Definition _s572_ (_s573_ : string)
-: option string :=
-
- let _s574_ := _s573_ in
- if ((string_startswith _s574_ "and")) then
- match (string_drop _s574_ (projT1 (string_length "and"))) with | s_ => Some (s_) end
+Definition _s833_ (_s834_ : string) : option string :=
+ let _s835_ := _s834_ in
+ if string_startswith _s835_ "and" then
+ match (string_drop _s835_ (projT1 (string_length "and"))) with | s_ => Some s_ end
else None.
-Definition _s568_ (_s569_ : string)
-: option string :=
-
- let _s570_ := _s569_ in
- if ((string_startswith _s570_ "sltu")) then
- match (string_drop _s570_ (projT1 (string_length "sltu"))) with | s_ => Some (s_) end
+Definition _s829_ (_s830_ : string) : option string :=
+ let _s831_ := _s830_ in
+ if string_startswith _s831_ "sltu" then
+ match (string_drop _s831_ (projT1 (string_length "sltu"))) with | s_ => Some s_ end
else None.
-Definition _s564_ (_s565_ : string)
-: option string :=
-
- let _s566_ := _s565_ in
- if ((string_startswith _s566_ "slt")) then
- match (string_drop _s566_ (projT1 (string_length "slt"))) with | s_ => Some (s_) end
+Definition _s825_ (_s826_ : string) : option string :=
+ let _s827_ := _s826_ in
+ if string_startswith _s827_ "slt" then
+ match (string_drop _s827_ (projT1 (string_length "slt"))) with | s_ => Some s_ end
else None.
-Definition _s560_ (_s561_ : string)
-: option string :=
-
- let _s562_ := _s561_ in
- if ((string_startswith _s562_ "add")) then
- match (string_drop _s562_ (projT1 (string_length "add"))) with | s_ => Some (s_) end
- else None.
-
-Definition rtype_mnemonic_matches_prefix (arg_ : string)
-: M (option ((rop * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s563_ := arg_ in
- (if ((match (_s560_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s560_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_ADD, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s564_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s564_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLT, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s568_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s568_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLTU, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s572_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s572_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_AND, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s576_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s576_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_OR, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s580_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s580_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_XOR, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s584_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s584_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLL, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s588_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s588_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRL, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s592_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s592_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SUB, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s596_ _s563_) with | Some (s_) => true | _ => false end)) then
- (match (_s596_ _s563_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRA, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((rop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
+Definition _s821_ (_s822_ : string) : option string :=
+ let _s823_ := _s822_ in
+ if string_startswith _s823_ "add" then
+ match (string_drop _s823_ (projT1 (string_length "add"))) with | s_ => Some s_ end
+ else None.
+
+Definition rtype_mnemonic_matches_prefix (arg_ : string)
+: M (option ((rop * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s824_ := arg_ in
+ (if match (_s821_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s821_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_ADD, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((rop * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((rop * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((rop * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s825_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s825_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLT, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s829_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s829_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLTU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s833_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s833_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_AND, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s837_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s837_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_OR, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s841_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s841_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_XOR, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s845_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s845_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLL, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s849_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s849_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRL, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s853_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s853_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SUB, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s857_ _s824_) with | Some s_ => true | _ => false end then
+ (match (_s857_ _s824_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRA, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((rop * {n : Z & ArithFact (n >=? 0)}))).
Definition extend_value {n : Z} (is_unsigned : bool) (value : MemoryOpResult (mword (8 * n)))
-`{ArithFact (0 < n /\ n <= 8)}
+`{ArithFact ((0 <? n) && (n <=? 8))}
: MemoryOpResult (mword 64) :=
-
match value with
- | MemValue (v) =>
- MemValue (if sumbool_of_bool (is_unsigned) then EXTZ 64 v else (EXTS 64 v) : xlenbits)
- | MemException (e) => MemException (e)
+ | MemValue v =>
+ MemValue (if sumbool_of_bool is_unsigned then EXTZ 64 v else (EXTS 64 v) : xlenbits)
+ | MemException e => MemException e
end.
Definition process_load {n : Z}
-(rd : mword 5) (addr : mword 64) (value : MemoryOpResult (mword (8 * n))) (is_unsigned : bool)
-`{ArithFact (0 < n /\ n <= 8)}
+(rd : mword 5) (vaddr : mword 64) (value : MemoryOpResult (mword (8 * n))) (is_unsigned : bool)
+`{ArithFact ((0 <? n) && (n <=? 8))}
: M (Retired) :=
-
(match (extend_value is_unsigned value) with
- | MemValue (result) =>
- (wX (projT1 (regidx_to_regno rd)) result) >> returnm (RETIRE_SUCCESS : Retired)
- | MemException (e) => (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
+ | MemValue result => (wX_bits rd result) >> returnm RETIRE_SUCCESS
+ | MemException e => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
end)
: M (Retired).
-Definition check_misaligned (vaddr : mword 64) (width : word_width)
-: M (bool) :=
-
- (if ((plat_enable_misaligned_access tt)) then
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
- else
- (match width with
- | BYTE =>
- returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))}))
- | HALF =>
- (bit_to_bool (access_vec_dec vaddr 0)) >>= fun w__0 : bool =>
- returnm ((Bool.eqb w__0 true)
- : bool)
- | WORD =>
- (or_boolM
- ((bit_to_bool (access_vec_dec vaddr 0)) >>= fun w__1 : bool =>
- returnm ((Bool.eqb w__1 true)
- : bool))
- ((bit_to_bool (access_vec_dec vaddr 1)) >>= fun w__2 : bool =>
- returnm ((Bool.eqb w__2 true)
- : bool)))
- : M (bool)
- | DOUBLE =>
- (or_boolM
- ((bit_to_bool (access_vec_dec vaddr 0)) >>= fun w__4 : bool =>
- returnm ((Bool.eqb w__4 true)
- : bool))
- ((or_boolM
- ((bit_to_bool (access_vec_dec vaddr 1)) >>= fun w__5 : bool =>
- returnm ((Bool.eqb w__5 true)
- : bool))
- ((bit_to_bool (access_vec_dec vaddr 2)) >>= fun w__6 : bool =>
- returnm ((Bool.eqb w__6 true)
- : bool)))
- : M (bool)))
- : M (bool)
- end)
- : M (bool))
- : M (bool).
-
-Definition maybe_aq_forwards (arg_ : bool)
-: string :=
-
+Definition check_misaligned (vaddr : mword 64) (width : word_width) : bool :=
+ if plat_enable_misaligned_access tt then false
+ else
+ match width with
+ | BYTE => false
+ | HALF => eq_bit (access_vec_dec vaddr 0) B1
+ | WORD => orb (eq_bit (access_vec_dec vaddr 0) B1) (eq_bit (access_vec_dec vaddr 1) B1)
+ | DOUBLE =>
+ orb (eq_bit (access_vec_dec vaddr 0) B1)
+ (orb (eq_bit (access_vec_dec vaddr 1) B1) (eq_bit (access_vec_dec vaddr 2) B1))
+ end.
+
+Definition maybe_aq_forwards (arg_ : bool) : string :=
match arg_ with | true => ".aq" | false => "" end.
-Definition maybe_aq_backwards (arg_ : string)
-: M (bool) :=
-
+Definition maybe_aq_backwards (arg_ : string) : M (bool) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ ".aq")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq p0_ "")) then
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ (if generic_eq p0_ ".aq" then returnm true
+ else if generic_eq p0_ "" then returnm false
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bool).
-Definition maybe_aq_forwards_matches (arg_ : bool)
-: bool :=
-
+Definition maybe_aq_forwards_matches (arg_ : bool) : bool :=
match arg_ with | true => true | false => true end.
-Definition maybe_aq_backwards_matches (arg_ : string)
-: bool :=
-
+Definition maybe_aq_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ ".aq")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ ".aq" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s604_ (_s605_ : string)
-: option string :=
-
- let _s606_ := _s605_ in
- if ((string_startswith _s606_ "")) then
- match (string_drop _s606_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s865_ (_s866_ : string) : option string :=
+ let _s867_ := _s866_ in
+ if string_startswith _s867_ "" then
+ match (string_drop _s867_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s600_ (_s601_ : string)
-: option string :=
-
- let _s602_ := _s601_ in
- if ((string_startswith _s602_ ".aq")) then
- match (string_drop _s602_ (projT1 (string_length ".aq"))) with | s_ => Some (s_) end
- else None.
-
-Definition maybe_aq_matches_prefix (arg_ : string)
-: M (option ((bool * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s603_ := arg_ in
- (if ((match (_s600_ _s603_) with | Some (s_) => true | _ => false end)) then
- (match (_s600_ _s603_) with
- | Some (s_) =>
- returnm ((Some
- ((true, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s604_ _s603_) with | Some (s_) => true | _ => false end)) then
- (match (_s604_ _s603_) with
- | Some (s_) =>
- returnm ((Some
- ((false, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
+Definition _s861_ (_s862_ : string) : option string :=
+ let _s863_ := _s862_ in
+ if string_startswith _s863_ ".aq" then
+ match (string_drop _s863_ (projT1 (string_length ".aq"))) with | s_ => Some s_ end
+ else None.
+
+Definition maybe_aq_matches_prefix (arg_ : string)
+: M (option ((bool * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s864_ := arg_ in
+ (if match (_s861_ _s864_) with | Some s_ => true | _ => false end then
+ (match (_s861_ _s864_) with
+ | Some s_ =>
+ returnm (Some
+ (true, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((bool * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((bool * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s865_ _s864_) with | Some s_ => true | _ => false end then
+ (match (_s865_ _s864_) with
+ | Some s_ =>
+ returnm (Some
+ (false, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)}))).
-Definition maybe_rl_forwards (arg_ : bool)
-: string :=
-
+Definition maybe_rl_forwards (arg_ : bool) : string :=
match arg_ with | true => ".rl" | false => "" end.
-Definition maybe_rl_backwards (arg_ : string)
-: M (bool) :=
-
+Definition maybe_rl_backwards (arg_ : string) : M (bool) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ ".rl")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq p0_ "")) then
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ (if generic_eq p0_ ".rl" then returnm true
+ else if generic_eq p0_ "" then returnm false
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bool).
-Definition maybe_rl_forwards_matches (arg_ : bool)
-: bool :=
-
+Definition maybe_rl_forwards_matches (arg_ : bool) : bool :=
match arg_ with | true => true | false => true end.
-Definition maybe_rl_backwards_matches (arg_ : string)
-: bool :=
-
+Definition maybe_rl_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ ".rl")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ ".rl" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s612_ (_s613_ : string)
-: option string :=
-
- let _s614_ := _s613_ in
- if ((string_startswith _s614_ "")) then
- match (string_drop _s614_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s873_ (_s874_ : string) : option string :=
+ let _s875_ := _s874_ in
+ if string_startswith _s875_ "" then
+ match (string_drop _s875_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s608_ (_s609_ : string)
-: option string :=
-
- let _s610_ := _s609_ in
- if ((string_startswith _s610_ ".rl")) then
- match (string_drop _s610_ (projT1 (string_length ".rl"))) with | s_ => Some (s_) end
- else None.
-
-Definition maybe_rl_matches_prefix (arg_ : string)
-: M (option ((bool * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s611_ := arg_ in
- (if ((match (_s608_ _s611_) with | Some (s_) => true | _ => false end)) then
- (match (_s608_ _s611_) with
- | Some (s_) =>
- returnm ((Some
- ((true, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s612_ _s611_) with | Some (s_) => true | _ => false end)) then
- (match (_s612_ _s611_) with
- | Some (s_) =>
- returnm ((Some
- ((false, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
+Definition _s869_ (_s870_ : string) : option string :=
+ let _s871_ := _s870_ in
+ if string_startswith _s871_ ".rl" then
+ match (string_drop _s871_ (projT1 (string_length ".rl"))) with | s_ => Some s_ end
+ else None.
+
+Definition maybe_rl_matches_prefix (arg_ : string)
+: M (option ((bool * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s872_ := arg_ in
+ (if match (_s869_ _s872_) with | Some s_ => true | _ => false end then
+ (match (_s869_ _s872_) with
+ | Some s_ =>
+ returnm (Some
+ (true, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((bool * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((bool * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s873_ _s872_) with | Some s_ => true | _ => false end then
+ (match (_s873_ _s872_) with
+ | Some s_ =>
+ returnm (Some
+ (false, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)}))).
-Definition maybe_u_forwards (arg_ : bool)
-: string :=
-
+Definition maybe_u_forwards (arg_ : bool) : string :=
match arg_ with | true => "u" | false => "" end.
-Definition maybe_u_backwards (arg_ : string)
-: M (bool) :=
-
+Definition maybe_u_backwards (arg_ : string) : M (bool) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "u")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq p0_ "")) then
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ (if generic_eq p0_ "u" then returnm true
+ else if generic_eq p0_ "" then returnm false
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bool).
-Definition maybe_u_forwards_matches (arg_ : bool)
-: bool :=
-
+Definition maybe_u_forwards_matches (arg_ : bool) : bool :=
match arg_ with | true => true | false => true end.
-Definition maybe_u_backwards_matches (arg_ : string)
-: bool :=
-
+Definition maybe_u_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "u")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ "u" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s620_ (_s621_ : string)
-: option string :=
-
- let _s622_ := _s621_ in
- if ((string_startswith _s622_ "")) then
- match (string_drop _s622_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s881_ (_s882_ : string) : option string :=
+ let _s883_ := _s882_ in
+ if string_startswith _s883_ "" then
+ match (string_drop _s883_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s616_ (_s617_ : string)
-: option string :=
-
- let _s618_ := _s617_ in
- if ((string_startswith _s618_ "u")) then
- match (string_drop _s618_ (projT1 (string_length "u"))) with | s_ => Some (s_) end
- else None.
-
-Definition maybe_u_matches_prefix (arg_ : string)
-: M (option ((bool * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s619_ := arg_ in
- (if ((match (_s616_ _s619_) with | Some (s_) => true | _ => false end)) then
- (match (_s616_ _s619_) with
- | Some (s_) =>
- returnm ((Some
- ((true, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s620_ _s619_) with | Some (s_) => true | _ => false end)) then
- (match (_s620_ _s619_) with
- | Some (s_) =>
- returnm ((Some
- ((false, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
+Definition _s877_ (_s878_ : string) : option string :=
+ let _s879_ := _s878_ in
+ if string_startswith _s879_ "u" then
+ match (string_drop _s879_ (projT1 (string_length "u"))) with | s_ => Some s_ end
+ else None.
+
+Definition maybe_u_matches_prefix (arg_ : string)
+: M (option ((bool * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s880_ := arg_ in
+ (if match (_s877_ _s880_) with | Some s_ => true | _ => false end then
+ (match (_s877_ _s880_) with
+ | Some s_ =>
+ returnm (Some
+ (true, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((bool * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((bool * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s881_ _s880_) with | Some s_ => true | _ => false end then
+ (match (_s881_ _s880_) with
+ | Some s_ =>
+ returnm (Some
+ (false, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)}))).
-Definition shiftw_mnemonic_forwards (arg_ : sop)
-: string :=
-
+Definition shiftw_mnemonic_forwards (arg_ : sop) : string :=
match arg_ with | RISCV_SLLI => "slli" | RISCV_SRLI => "srli" | RISCV_SRAI => "srai" end.
-Definition shiftw_mnemonic_backwards (arg_ : string)
-: M (sop) :=
-
+Definition shiftw_mnemonic_backwards (arg_ : string) : M (sop) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "slli")) then returnm (RISCV_SLLI : sop)
- else if ((generic_eq p0_ "srli")) then returnm (RISCV_SRLI : sop)
- else if ((generic_eq p0_ "srai")) then returnm (RISCV_SRAI : sop)
+ (if generic_eq p0_ "slli" then returnm RISCV_SLLI
+ else if generic_eq p0_ "srli" then returnm RISCV_SRLI
+ else if generic_eq p0_ "srai" then returnm RISCV_SRAI
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (sop).
-Definition shiftw_mnemonic_forwards_matches (arg_ : sop)
-: bool :=
-
+Definition shiftw_mnemonic_forwards_matches (arg_ : sop) : bool :=
match arg_ with | RISCV_SLLI => true | RISCV_SRLI => true | RISCV_SRAI => true end.
-Definition shiftw_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition shiftw_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "slli")) then true
- else if ((generic_eq p0_ "srli")) then true
- else if ((generic_eq p0_ "srai")) then true
+ if generic_eq p0_ "slli" then true
+ else if generic_eq p0_ "srli" then true
+ else if generic_eq p0_ "srai" then true
else false.
-Definition _s632_ (_s633_ : string)
-: option string :=
-
- let _s634_ := _s633_ in
- if ((string_startswith _s634_ "srai")) then
- match (string_drop _s634_ (projT1 (string_length "srai"))) with | s_ => Some (s_) end
+Definition _s893_ (_s894_ : string) : option string :=
+ let _s895_ := _s894_ in
+ if string_startswith _s895_ "srai" then
+ match (string_drop _s895_ (projT1 (string_length "srai"))) with | s_ => Some s_ end
else None.
-Definition _s628_ (_s629_ : string)
-: option string :=
-
- let _s630_ := _s629_ in
- if ((string_startswith _s630_ "srli")) then
- match (string_drop _s630_ (projT1 (string_length "srli"))) with | s_ => Some (s_) end
+Definition _s889_ (_s890_ : string) : option string :=
+ let _s891_ := _s890_ in
+ if string_startswith _s891_ "srli" then
+ match (string_drop _s891_ (projT1 (string_length "srli"))) with | s_ => Some s_ end
else None.
-Definition _s624_ (_s625_ : string)
-: option string :=
-
- let _s626_ := _s625_ in
- if ((string_startswith _s626_ "slli")) then
- match (string_drop _s626_ (projT1 (string_length "slli"))) with | s_ => Some (s_) end
- else None.
-
-Definition shiftw_mnemonic_matches_prefix (arg_ : string)
-: M (option ((sop * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s627_ := arg_ in
- (if ((match (_s624_ _s627_) with | Some (s_) => true | _ => false end)) then
- (match (_s624_ _s627_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLLI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s628_ _s627_) with | Some (s_) => true | _ => false end)) then
- (match (_s628_ _s627_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRLI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s632_ _s627_) with | Some (s_) => true | _ => false end)) then
- (match (_s632_ _s627_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRAI, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
+Definition _s885_ (_s886_ : string) : option string :=
+ let _s887_ := _s886_ in
+ if string_startswith _s887_ "slli" then
+ match (string_drop _s887_ (projT1 (string_length "slli"))) with | s_ => Some s_ end
+ else None.
+
+Definition shiftw_mnemonic_matches_prefix (arg_ : string)
+: M (option ((sop * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s888_ := arg_ in
+ (if match (_s885_ _s888_) with | Some s_ => true | _ => false end then
+ (match (_s885_ _s888_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLLI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((sop * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((sop * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((sop * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s889_ _s888_) with | Some s_ => true | _ => false end then
+ (match (_s889_ _s888_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRLI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s893_ _s888_) with | Some s_ => true | _ => false end then
+ (match (_s893_ _s888_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRAI, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((sop * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((sop * {n : Z & ArithFact (n >=? 0)}))).
-Definition rtypew_mnemonic_forwards (arg_ : ropw)
-: string :=
-
+Definition rtypew_mnemonic_forwards (arg_ : ropw) : string :=
match arg_ with
| RISCV_ADDW => "addw"
| RISCV_SUBW => "subw"
@@ -15146,21 +14390,17 @@ Definition rtypew_mnemonic_forwards (arg_ : ropw)
| RISCV_SRAW => "sraw"
end.
-Definition rtypew_mnemonic_backwards (arg_ : string)
-: M (ropw) :=
-
+Definition rtypew_mnemonic_backwards (arg_ : string) : M (ropw) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "addw")) then returnm (RISCV_ADDW : ropw)
- else if ((generic_eq p0_ "subw")) then returnm (RISCV_SUBW : ropw)
- else if ((generic_eq p0_ "sllw")) then returnm (RISCV_SLLW : ropw)
- else if ((generic_eq p0_ "srlw")) then returnm (RISCV_SRLW : ropw)
- else if ((generic_eq p0_ "sraw")) then returnm (RISCV_SRAW : ropw)
+ (if generic_eq p0_ "addw" then returnm RISCV_ADDW
+ else if generic_eq p0_ "subw" then returnm RISCV_SUBW
+ else if generic_eq p0_ "sllw" then returnm RISCV_SLLW
+ else if generic_eq p0_ "srlw" then returnm RISCV_SRLW
+ else if generic_eq p0_ "sraw" then returnm RISCV_SRAW
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (ropw).
-Definition rtypew_mnemonic_forwards_matches (arg_ : ropw)
-: bool :=
-
+Definition rtypew_mnemonic_forwards_matches (arg_ : ropw) : bool :=
match arg_ with
| RISCV_ADDW => true
| RISCV_SUBW => true
@@ -15169,531 +14409,441 @@ Definition rtypew_mnemonic_forwards_matches (arg_ : ropw)
| RISCV_SRAW => true
end.
-Definition rtypew_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition rtypew_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "addw")) then true
- else if ((generic_eq p0_ "subw")) then true
- else if ((generic_eq p0_ "sllw")) then true
- else if ((generic_eq p0_ "srlw")) then true
- else if ((generic_eq p0_ "sraw")) then true
+ if generic_eq p0_ "addw" then true
+ else if generic_eq p0_ "subw" then true
+ else if generic_eq p0_ "sllw" then true
+ else if generic_eq p0_ "srlw" then true
+ else if generic_eq p0_ "sraw" then true
else false.
-Definition _s652_ (_s653_ : string)
-: option string :=
-
- let _s654_ := _s653_ in
- if ((string_startswith _s654_ "sraw")) then
- match (string_drop _s654_ (projT1 (string_length "sraw"))) with | s_ => Some (s_) end
+Definition _s913_ (_s914_ : string) : option string :=
+ let _s915_ := _s914_ in
+ if string_startswith _s915_ "sraw" then
+ match (string_drop _s915_ (projT1 (string_length "sraw"))) with | s_ => Some s_ end
else None.
-Definition _s648_ (_s649_ : string)
-: option string :=
-
- let _s650_ := _s649_ in
- if ((string_startswith _s650_ "srlw")) then
- match (string_drop _s650_ (projT1 (string_length "srlw"))) with | s_ => Some (s_) end
+Definition _s909_ (_s910_ : string) : option string :=
+ let _s911_ := _s910_ in
+ if string_startswith _s911_ "srlw" then
+ match (string_drop _s911_ (projT1 (string_length "srlw"))) with | s_ => Some s_ end
else None.
-Definition _s644_ (_s645_ : string)
-: option string :=
-
- let _s646_ := _s645_ in
- if ((string_startswith _s646_ "sllw")) then
- match (string_drop _s646_ (projT1 (string_length "sllw"))) with | s_ => Some (s_) end
+Definition _s905_ (_s906_ : string) : option string :=
+ let _s907_ := _s906_ in
+ if string_startswith _s907_ "sllw" then
+ match (string_drop _s907_ (projT1 (string_length "sllw"))) with | s_ => Some s_ end
else None.
-Definition _s640_ (_s641_ : string)
-: option string :=
-
- let _s642_ := _s641_ in
- if ((string_startswith _s642_ "subw")) then
- match (string_drop _s642_ (projT1 (string_length "subw"))) with | s_ => Some (s_) end
+Definition _s901_ (_s902_ : string) : option string :=
+ let _s903_ := _s902_ in
+ if string_startswith _s903_ "subw" then
+ match (string_drop _s903_ (projT1 (string_length "subw"))) with | s_ => Some s_ end
else None.
-Definition _s636_ (_s637_ : string)
-: option string :=
-
- let _s638_ := _s637_ in
- if ((string_startswith _s638_ "addw")) then
- match (string_drop _s638_ (projT1 (string_length "addw"))) with | s_ => Some (s_) end
- else None.
-
-Definition rtypew_mnemonic_matches_prefix (arg_ : string)
-: M (option ((ropw * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s639_ := arg_ in
- (if ((match (_s636_ _s639_) with | Some (s_) => true | _ => false end)) then
- (match (_s636_ _s639_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_ADDW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ropw * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s640_ _s639_) with | Some (s_) => true | _ => false end)) then
- (match (_s640_ _s639_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SUBW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ropw * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s644_ _s639_) with | Some (s_) => true | _ => false end)) then
- (match (_s644_ _s639_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLLW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ropw * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s648_ _s639_) with | Some (s_) => true | _ => false end)) then
- (match (_s648_ _s639_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRLW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ropw * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s652_ _s639_) with | Some (s_) => true | _ => false end)) then
- (match (_s652_ _s639_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRAW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ropw * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
+Definition _s897_ (_s898_ : string) : option string :=
+ let _s899_ := _s898_ in
+ if string_startswith _s899_ "addw" then
+ match (string_drop _s899_ (projT1 (string_length "addw"))) with | s_ => Some s_ end
+ else None.
+
+Definition rtypew_mnemonic_matches_prefix (arg_ : string)
+: M (option ((ropw * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s900_ := arg_ in
+ (if match (_s897_ _s900_) with | Some s_ => true | _ => false end then
+ (match (_s897_ _s900_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_ADDW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s901_ _s900_) with | Some s_ => true | _ => false end then
+ (match (_s901_ _s900_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SUBW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s905_ _s900_) with | Some s_ => true | _ => false end then
+ (match (_s905_ _s900_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLLW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s909_ _s900_) with | Some s_ => true | _ => false end then
+ (match (_s909_ _s900_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRLW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ropw * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((ropw * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ropw * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s913_ _s900_) with | Some s_ => true | _ => false end then
+ (match (_s913_ _s900_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRAW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((ropw * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((ropw * {n : Z & ArithFact (n >=? 0)}))).
-Definition shiftiwop_mnemonic_forwards (arg_ : sopw)
-: string :=
-
+Definition shiftiwop_mnemonic_forwards (arg_ : sopw) : string :=
match arg_ with | RISCV_SLLIW => "slliw" | RISCV_SRLIW => "srliw" | RISCV_SRAIW => "sraiw" end.
-Definition shiftiwop_mnemonic_backwards (arg_ : string)
-: M (sopw) :=
-
+Definition shiftiwop_mnemonic_backwards (arg_ : string) : M (sopw) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "slliw")) then returnm (RISCV_SLLIW : sopw)
- else if ((generic_eq p0_ "srliw")) then returnm (RISCV_SRLIW : sopw)
- else if ((generic_eq p0_ "sraiw")) then returnm (RISCV_SRAIW : sopw)
+ (if generic_eq p0_ "slliw" then returnm RISCV_SLLIW
+ else if generic_eq p0_ "srliw" then returnm RISCV_SRLIW
+ else if generic_eq p0_ "sraiw" then returnm RISCV_SRAIW
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (sopw).
-Definition shiftiwop_mnemonic_forwards_matches (arg_ : sopw)
-: bool :=
-
+Definition shiftiwop_mnemonic_forwards_matches (arg_ : sopw) : bool :=
match arg_ with | RISCV_SLLIW => true | RISCV_SRLIW => true | RISCV_SRAIW => true end.
-Definition shiftiwop_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition shiftiwop_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "slliw")) then true
- else if ((generic_eq p0_ "srliw")) then true
- else if ((generic_eq p0_ "sraiw")) then true
+ if generic_eq p0_ "slliw" then true
+ else if generic_eq p0_ "srliw" then true
+ else if generic_eq p0_ "sraiw" then true
else false.
-Definition _s664_ (_s665_ : string)
-: option string :=
-
- let _s666_ := _s665_ in
- if ((string_startswith _s666_ "sraiw")) then
- match (string_drop _s666_ (projT1 (string_length "sraiw"))) with | s_ => Some (s_) end
+Definition _s925_ (_s926_ : string) : option string :=
+ let _s927_ := _s926_ in
+ if string_startswith _s927_ "sraiw" then
+ match (string_drop _s927_ (projT1 (string_length "sraiw"))) with | s_ => Some s_ end
else None.
-Definition _s660_ (_s661_ : string)
-: option string :=
-
- let _s662_ := _s661_ in
- if ((string_startswith _s662_ "srliw")) then
- match (string_drop _s662_ (projT1 (string_length "srliw"))) with | s_ => Some (s_) end
+Definition _s921_ (_s922_ : string) : option string :=
+ let _s923_ := _s922_ in
+ if string_startswith _s923_ "srliw" then
+ match (string_drop _s923_ (projT1 (string_length "srliw"))) with | s_ => Some s_ end
else None.
-Definition _s656_ (_s657_ : string)
-: option string :=
-
- let _s658_ := _s657_ in
- if ((string_startswith _s658_ "slliw")) then
- match (string_drop _s658_ (projT1 (string_length "slliw"))) with | s_ => Some (s_) end
- else None.
-
-Definition shiftiwop_mnemonic_matches_prefix (arg_ : string)
-: M (option ((sopw * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s659_ := arg_ in
- (if ((match (_s656_ _s659_) with | Some (s_) => true | _ => false end)) then
- (match (_s656_ _s659_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SLLIW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sopw * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sopw * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((sopw * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s660_ _s659_) with | Some (s_) => true | _ => false end)) then
- (match (_s660_ _s659_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRLIW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sopw * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sopw * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((sopw * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s664_ _s659_) with | Some (s_) => true | _ => false end)) then
- (match (_s664_ _s659_) with
- | Some (s_) =>
- returnm ((Some
- ((RISCV_SRAIW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((sopw * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((sopw * {n : Z & ArithFact (n >= 0)})))
+Definition _s917_ (_s918_ : string) : option string :=
+ let _s919_ := _s918_ in
+ if string_startswith _s919_ "slliw" then
+ match (string_drop _s919_ (projT1 (string_length "slliw"))) with | s_ => Some s_ end
+ else None.
+
+Definition shiftiwop_mnemonic_matches_prefix (arg_ : string)
+: M (option ((sopw * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s920_ := arg_ in
+ (if match (_s917_ _s920_) with | Some s_ => true | _ => false end then
+ (match (_s917_ _s920_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SLLIW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sopw * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((sopw * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s921_ _s920_) with | Some s_ => true | _ => false end then
+ (match (_s921_ _s920_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRLIW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sopw * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((sopw * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((sopw * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((sopw * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((sopw * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s925_ _s920_) with | Some s_ => true | _ => false end then
+ (match (_s925_ _s920_) with
+ | Some s_ =>
+ returnm (Some
+ (RISCV_SRAIW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((sopw * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((sopw * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((sopw * {n : Z & ArithFact (n >=? 0)}))).
-Definition bit_maybe_r_forwards (arg_ : mword 1)
-: M (string) :=
-
+Definition bit_maybe_r_forwards (arg_ : mword 1) : M (string) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then returnm ("r" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then returnm ("" : string)
+ (if eq_vec b__0 ('b"1" : mword 1) then returnm "r"
+ else if eq_vec b__0 ('b"0" : mword 1) then returnm ""
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string).
-Definition bit_maybe_r_backwards (arg_ : string)
-: M (mword 1) :=
-
+Definition bit_maybe_r_backwards (arg_ : string) : M (mword 1) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "r")) then returnm ((vec_of_bits [B1] : mword 1) : mword 1)
- else if ((generic_eq p0_ "")) then returnm ((vec_of_bits [B0] : mword 1) : mword 1)
+ (if generic_eq p0_ "r" then returnm ('b"1" : mword 1)
+ else if generic_eq p0_ "" then returnm ('b"0" : mword 1)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 1).
-Definition bit_maybe_r_forwards_matches (arg_ : mword 1)
-: bool :=
-
+Definition bit_maybe_r_forwards_matches (arg_ : mword 1) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
+ if eq_vec b__0 ('b"1" : mword 1) then true
+ else if eq_vec b__0 ('b"0" : mword 1) then true
else false.
-Definition bit_maybe_r_backwards_matches (arg_ : string)
-: bool :=
-
+Definition bit_maybe_r_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "r")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ "r" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s672_ (_s673_ : string)
-: option string :=
-
- let _s674_ := _s673_ in
- if ((string_startswith _s674_ "")) then
- match (string_drop _s674_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s933_ (_s934_ : string) : option string :=
+ let _s935_ := _s934_ in
+ if string_startswith _s935_ "" then
+ match (string_drop _s935_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s668_ (_s669_ : string)
-: option string :=
-
- let _s670_ := _s669_ in
- if ((string_startswith _s670_ "r")) then
- match (string_drop _s670_ (projT1 (string_length "r"))) with | s_ => Some (s_) end
- else None.
-
-Definition bit_maybe_r_matches_prefix (arg_ : string)
-: M (option ((mword 1 * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s671_ := arg_ in
- (if ((match (_s668_ _s671_) with | Some (s_) => true | _ => false end)) then
- (match (_s668_ _s671_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1] : mword 1), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s672_ _s671_) with | Some (s_) => true | _ => false end)) then
- (match (_s672_ _s671_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0] : mword 1), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
+Definition _s929_ (_s930_ : string) : option string :=
+ let _s931_ := _s930_ in
+ if string_startswith _s931_ "r" then
+ match (string_drop _s931_ (projT1 (string_length "r"))) with | s_ => Some s_ end
+ else None.
+
+Definition bit_maybe_r_matches_prefix (arg_ : string)
+: M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s932_ := arg_ in
+ (if match (_s929_ _s932_) with | Some s_ => true | _ => false end then
+ (match (_s929_ _s932_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"1"
+ : mword 1, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((mword 1 * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s933_ _s932_) with | Some s_ => true | _ => false end then
+ (match (_s933_ _s932_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"0"
+ : mword 1, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)}))).
-Definition bit_maybe_w_forwards (arg_ : mword 1)
-: M (string) :=
-
+Definition bit_maybe_w_forwards (arg_ : mword 1) : M (string) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then returnm ("w" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then returnm ("" : string)
+ (if eq_vec b__0 ('b"1" : mword 1) then returnm "w"
+ else if eq_vec b__0 ('b"0" : mword 1) then returnm ""
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string).
-Definition bit_maybe_w_backwards (arg_ : string)
-: M (mword 1) :=
-
+Definition bit_maybe_w_backwards (arg_ : string) : M (mword 1) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "w")) then returnm ((vec_of_bits [B1] : mword 1) : mword 1)
- else if ((generic_eq p0_ "")) then returnm ((vec_of_bits [B0] : mword 1) : mword 1)
+ (if generic_eq p0_ "w" then returnm ('b"1" : mword 1)
+ else if generic_eq p0_ "" then returnm ('b"0" : mword 1)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 1).
-Definition bit_maybe_w_forwards_matches (arg_ : mword 1)
-: bool :=
-
+Definition bit_maybe_w_forwards_matches (arg_ : mword 1) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
+ if eq_vec b__0 ('b"1" : mword 1) then true
+ else if eq_vec b__0 ('b"0" : mword 1) then true
else false.
-Definition bit_maybe_w_backwards_matches (arg_ : string)
-: bool :=
-
+Definition bit_maybe_w_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "w")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ "w" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s680_ (_s681_ : string)
-: option string :=
-
- let _s682_ := _s681_ in
- if ((string_startswith _s682_ "")) then
- match (string_drop _s682_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s941_ (_s942_ : string) : option string :=
+ let _s943_ := _s942_ in
+ if string_startswith _s943_ "" then
+ match (string_drop _s943_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s676_ (_s677_ : string)
-: option string :=
-
- let _s678_ := _s677_ in
- if ((string_startswith _s678_ "w")) then
- match (string_drop _s678_ (projT1 (string_length "w"))) with | s_ => Some (s_) end
- else None.
-
-Definition bit_maybe_w_matches_prefix (arg_ : string)
-: M (option ((mword 1 * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s679_ := arg_ in
- (if ((match (_s676_ _s679_) with | Some (s_) => true | _ => false end)) then
- (match (_s676_ _s679_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1] : mword 1), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s680_ _s679_) with | Some (s_) => true | _ => false end)) then
- (match (_s680_ _s679_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0] : mword 1), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
+Definition _s937_ (_s938_ : string) : option string :=
+ let _s939_ := _s938_ in
+ if string_startswith _s939_ "w" then
+ match (string_drop _s939_ (projT1 (string_length "w"))) with | s_ => Some s_ end
+ else None.
+
+Definition bit_maybe_w_matches_prefix (arg_ : string)
+: M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s940_ := arg_ in
+ (if match (_s937_ _s940_) with | Some s_ => true | _ => false end then
+ (match (_s937_ _s940_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"1"
+ : mword 1, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((mword 1 * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s941_ _s940_) with | Some s_ => true | _ => false end then
+ (match (_s941_ _s940_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"0"
+ : mword 1, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)}))).
-Definition bit_maybe_i_forwards (arg_ : mword 1)
-: M (string) :=
-
+Definition bit_maybe_i_forwards (arg_ : mword 1) : M (string) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then returnm ("i" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then returnm ("" : string)
+ (if eq_vec b__0 ('b"1" : mword 1) then returnm "i"
+ else if eq_vec b__0 ('b"0" : mword 1) then returnm ""
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string).
-Definition bit_maybe_i_backwards (arg_ : string)
-: M (mword 1) :=
-
+Definition bit_maybe_i_backwards (arg_ : string) : M (mword 1) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "i")) then returnm ((vec_of_bits [B1] : mword 1) : mword 1)
- else if ((generic_eq p0_ "")) then returnm ((vec_of_bits [B0] : mword 1) : mword 1)
+ (if generic_eq p0_ "i" then returnm ('b"1" : mword 1)
+ else if generic_eq p0_ "" then returnm ('b"0" : mword 1)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 1).
-Definition bit_maybe_i_forwards_matches (arg_ : mword 1)
-: bool :=
-
+Definition bit_maybe_i_forwards_matches (arg_ : mword 1) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
+ if eq_vec b__0 ('b"1" : mword 1) then true
+ else if eq_vec b__0 ('b"0" : mword 1) then true
else false.
-Definition bit_maybe_i_backwards_matches (arg_ : string)
-: bool :=
-
+Definition bit_maybe_i_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "i")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ "i" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s688_ (_s689_ : string)
-: option string :=
-
- let _s690_ := _s689_ in
- if ((string_startswith _s690_ "")) then
- match (string_drop _s690_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s949_ (_s950_ : string) : option string :=
+ let _s951_ := _s950_ in
+ if string_startswith _s951_ "" then
+ match (string_drop _s951_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s684_ (_s685_ : string)
-: option string :=
-
- let _s686_ := _s685_ in
- if ((string_startswith _s686_ "i")) then
- match (string_drop _s686_ (projT1 (string_length "i"))) with | s_ => Some (s_) end
- else None.
-
-Definition bit_maybe_i_matches_prefix (arg_ : string)
-: M (option ((mword 1 * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s687_ := arg_ in
- (if ((match (_s684_ _s687_) with | Some (s_) => true | _ => false end)) then
- (match (_s684_ _s687_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1] : mword 1), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s688_ _s687_) with | Some (s_) => true | _ => false end)) then
- (match (_s688_ _s687_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0] : mword 1), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
+Definition _s945_ (_s946_ : string) : option string :=
+ let _s947_ := _s946_ in
+ if string_startswith _s947_ "i" then
+ match (string_drop _s947_ (projT1 (string_length "i"))) with | s_ => Some s_ end
+ else None.
+
+Definition bit_maybe_i_matches_prefix (arg_ : string)
+: M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s948_ := arg_ in
+ (if match (_s945_ _s948_) with | Some s_ => true | _ => false end then
+ (match (_s945_ _s948_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"1"
+ : mword 1, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((mword 1 * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s949_ _s948_) with | Some s_ => true | _ => false end then
+ (match (_s949_ _s948_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"0"
+ : mword 1, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)}))).
-Definition bit_maybe_o_forwards (arg_ : mword 1)
-: M (string) :=
-
+Definition bit_maybe_o_forwards (arg_ : mword 1) : M (string) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then returnm ("o" : string)
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then returnm ("" : string)
+ (if eq_vec b__0 ('b"1" : mword 1) then returnm "o"
+ else if eq_vec b__0 ('b"0" : mword 1) then returnm ""
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string).
-Definition bit_maybe_o_backwards (arg_ : string)
-: M (mword 1) :=
-
+Definition bit_maybe_o_backwards (arg_ : string) : M (mword 1) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "o")) then returnm ((vec_of_bits [B1] : mword 1) : mword 1)
- else if ((generic_eq p0_ "")) then returnm ((vec_of_bits [B0] : mword 1) : mword 1)
+ (if generic_eq p0_ "o" then returnm ('b"1" : mword 1)
+ else if generic_eq p0_ "" then returnm ('b"0" : mword 1)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 1).
-Definition bit_maybe_o_forwards_matches (arg_ : mword 1)
-: bool :=
-
+Definition bit_maybe_o_forwards_matches (arg_ : mword 1) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B1] : mword 1))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0] : mword 1))) then true
+ if eq_vec b__0 ('b"1" : mword 1) then true
+ else if eq_vec b__0 ('b"0" : mword 1) then true
else false.
-Definition bit_maybe_o_backwards_matches (arg_ : string)
-: bool :=
-
+Definition bit_maybe_o_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "o")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ "o" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s696_ (_s697_ : string)
-: option string :=
-
- let _s698_ := _s697_ in
- if ((string_startswith _s698_ "")) then
- match (string_drop _s698_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s957_ (_s958_ : string) : option string :=
+ let _s959_ := _s958_ in
+ if string_startswith _s959_ "" then
+ match (string_drop _s959_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s692_ (_s693_ : string)
-: option string :=
-
- let _s694_ := _s693_ in
- if ((string_startswith _s694_ "o")) then
- match (string_drop _s694_ (projT1 (string_length "o"))) with | s_ => Some (s_) end
- else None.
-
-Definition bit_maybe_o_matches_prefix (arg_ : string)
-: M (option ((mword 1 * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s695_ := arg_ in
- (if ((match (_s692_ _s695_) with | Some (s_) => true | _ => false end)) then
- (match (_s692_ _s695_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B1] : mword 1), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s696_ _s695_) with | Some (s_) => true | _ => false end)) then
- (match (_s696_ _s695_) with
- | Some (s_) =>
- returnm ((Some
- (((vec_of_bits [B0] : mword 1), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
+Definition _s953_ (_s954_ : string) : option string :=
+ let _s955_ := _s954_ in
+ if string_startswith _s955_ "o" then
+ match (string_drop _s955_ (projT1 (string_length "o"))) with | s_ => Some s_ end
+ else None.
+
+Definition bit_maybe_o_matches_prefix (arg_ : string)
+: M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s956_ := arg_ in
+ (if match (_s953_ _s956_) with | Some s_ => true | _ => false end then
+ (match (_s953_ _s956_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"1"
+ : mword 1, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((mword 1 * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((mword 1 * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s957_ _s956_) with | Some s_ => true | _ => false end then
+ (match (_s957_ _s956_) with
+ | Some s_ =>
+ returnm (Some
+ ('b"0"
+ : mword 1, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 1 * {n : Z & ArithFact (n >=? 0)}))).
-Definition fence_bits_forwards (arg_ : mword 4)
-: M (string) :=
-
+Definition fence_bits_forwards (arg_ : mword 4) : M (string) :=
(match arg_ with
| v__0 =>
let i : bits 1 := subrange_vec_dec v__0 3 3 in
@@ -15705,230 +14855,202 @@ Definition fence_bits_forwards (arg_ : mword 4)
(bit_maybe_o_forwards o) >>= fun w__1 : string =>
(bit_maybe_r_forwards r) >>= fun w__2 : string =>
(bit_maybe_w_forwards w) >>= fun w__3 : string =>
- returnm ((string_append w__0
- (string_append w__1 (string_append w__2 (string_append w__3 ""))))
- : string)
+ returnm (string_append w__0 (string_append w__1 (string_append w__2 (string_append w__3 ""))))
end)
: M (string).
-Definition _s700_ (_s701_ : string)
-: M (option ((mword 1 * mword 1 * mword 1 * mword 1))) :=
-
- (match _s701_ with
- | _s702_ =>
- (bit_maybe_i_matches_prefix _s702_) >>= fun w__0 : option ((mword 1 * {n : Z & ArithFact (n >=
+Definition _s961_ (_s962_ : string) : M (option ((mword 1 * mword 1 * mword 1 * mword 1))) :=
+ (match _s962_ with
+ | _s963_ =>
+ (bit_maybe_i_matches_prefix _s963_) >>= fun w__0 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((i, existT _ _s703_ _)) =>
- (match (string_drop _s702_ _s703_) with
- | _s704_ =>
- (bit_maybe_o_matches_prefix _s704_) >>= fun w__1 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (i, existT _ _s964_ _) =>
+ (match (string_drop _s963_ _s964_) with
+ | _s965_ =>
+ (bit_maybe_o_matches_prefix _s965_) >>= fun w__1 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((o, existT _ _s705_ _)) =>
- (match (string_drop _s704_ _s705_) with
- | _s706_ =>
- (bit_maybe_r_matches_prefix _s706_) >>= fun w__2 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (o, existT _ _s966_ _) =>
+ (match (string_drop _s965_ _s966_) with
+ | _s967_ =>
+ (bit_maybe_r_matches_prefix _s967_) >>= fun w__2 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((r, existT _ _s707_ _)) =>
- (match (string_drop _s706_ _s707_) with
- | _s708_ =>
- (bit_maybe_w_matches_prefix _s708_) >>= fun w__3 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (r, existT _ _s968_ _) =>
+ (match (string_drop _s967_ _s968_) with
+ | _s969_ =>
+ (bit_maybe_w_matches_prefix _s969_) >>= fun w__3 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((w, existT _ _s709_ _)) =>
- let p0_ := string_drop _s708_ _s709_ in
- if ((generic_eq p0_ "")) then Some ((i, o, r, w))
- else None
- | _ => None
- end)
- : option ((mword 1 * mword 1 * mword 1 * mword 1)))
+ returnm (match w__3 with
+ | Some (w, existT _ _s970_ _) =>
+ let p0_ := string_drop _s969_ _s970_ in
+ if generic_eq p0_ "" then Some (i, o, r, w)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
- | _ => returnm (None : option ((mword 1 * mword 1 * mword 1 * mword 1)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
- | _ => returnm (None : option ((mword 1 * mword 1 * mword 1 * mword 1)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
- | _ => returnm (None : option ((mword 1 * mword 1 * mword 1 * mword 1)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1))).
-Definition fence_bits_backwards (arg_ : string)
-: M (mword 4) :=
-
- let _s710_ := arg_ in
- (_s700_ _s710_) >>= fun w__0 : option ((mword 1 * mword 1 * mword 1 * mword 1)) =>
- (if ((match w__0 with | Some ((i, o, r, w)) => true | _ => false end)) then
- (_s700_ _s710_) >>= fun w__1 : option ((mword 1 * mword 1 * mword 1 * mword 1)) =>
+Definition fence_bits_backwards (arg_ : string) : M (mword 4) :=
+ let _s971_ := arg_ in
+ (_s961_ _s971_) >>= fun w__0 : option ((mword 1 * mword 1 * mword 1 * mword 1)) =>
+ (if match w__0 with | Some (i, o, r, w) => true | _ => false end then
+ (_s961_ _s971_) >>= fun w__1 : option ((mword 1 * mword 1 * mword 1 * mword 1)) =>
(match w__1 with
- | Some ((i, o, r, w)) =>
- returnm ((concat_vec (i : bits 1)
- (concat_vec (o : bits 1) (concat_vec (r : bits 1) (w : bits 1))))
- : mword (1 + (1 + (1 + 1))))
+ | Some (i, o, r, w) =>
+ returnm (concat_vec (i : bits 1)
+ (concat_vec (o : bits 1) (concat_vec (r : bits 1) (w : bits 1))))
| _ => exit tt : M (mword 4)
end)
: M (mword 4)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 4).
-Definition fence_bits_forwards_matches (arg_ : mword 4)
-: bool :=
-
+Definition fence_bits_forwards_matches (arg_ : mword 4) : bool :=
match arg_ with | v__1 => true end.
-Definition _s711_ (_s712_ : string)
-: M (option ((mword 1 * mword 1 * mword 1 * mword 1))) :=
-
- (match _s712_ with
- | _s713_ =>
- (bit_maybe_i_matches_prefix _s713_) >>= fun w__0 : option ((mword 1 * {n : Z & ArithFact (n >=
+Definition _s972_ (_s973_ : string) : M (option ((mword 1 * mword 1 * mword 1 * mword 1))) :=
+ (match _s973_ with
+ | _s974_ =>
+ (bit_maybe_i_matches_prefix _s974_) >>= fun w__0 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((i, existT _ _s714_ _)) =>
- (match (string_drop _s713_ _s714_) with
- | _s715_ =>
- (bit_maybe_o_matches_prefix _s715_) >>= fun w__1 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (i, existT _ _s975_ _) =>
+ (match (string_drop _s974_ _s975_) with
+ | _s976_ =>
+ (bit_maybe_o_matches_prefix _s976_) >>= fun w__1 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((o, existT _ _s716_ _)) =>
- (match (string_drop _s715_ _s716_) with
- | _s717_ =>
- (bit_maybe_r_matches_prefix _s717_) >>= fun w__2 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (o, existT _ _s977_ _) =>
+ (match (string_drop _s976_ _s977_) with
+ | _s978_ =>
+ (bit_maybe_r_matches_prefix _s978_) >>= fun w__2 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((r, existT _ _s718_ _)) =>
- (match (string_drop _s717_ _s718_) with
- | _s719_ =>
- (bit_maybe_w_matches_prefix _s719_) >>= fun w__3 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (r, existT _ _s979_ _) =>
+ (match (string_drop _s978_ _s979_) with
+ | _s980_ =>
+ (bit_maybe_w_matches_prefix _s980_) >>= fun w__3 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((w, existT _ _s720_ _)) =>
- let p0_ := string_drop _s719_ _s720_ in
- if ((generic_eq p0_ "")) then Some ((i, o, r, w))
- else None
- | _ => None
- end)
- : option ((mword 1 * mword 1 * mword 1 * mword 1)))
+ returnm (match w__3 with
+ | Some (w, existT _ _s981_ _) =>
+ let p0_ := string_drop _s980_ _s981_ in
+ if generic_eq p0_ "" then Some (i, o, r, w)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
- | _ => returnm (None : option ((mword 1 * mword 1 * mword 1 * mword 1)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
- | _ => returnm (None : option ((mword 1 * mword 1 * mword 1 * mword 1)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
- | _ => returnm (None : option ((mword 1 * mword 1 * mword 1 * mword 1)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1))).
-Definition fence_bits_backwards_matches (arg_ : string)
-: M (bool) :=
-
- let _s721_ := arg_ in
- (_s711_ _s721_) >>= fun w__0 : option ((mword 1 * mword 1 * mword 1 * mword 1)) =>
- (if ((match w__0 with | Some ((i, o, r, w)) => true | _ => false end)) then
- (_s711_ _s721_) >>= fun w__1 : option ((mword 1 * mword 1 * mword 1 * mword 1)) =>
- (match w__1 with
- | Some ((i, o, r, w)) =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | _ => exit tt : M (bool)
- end)
+Definition fence_bits_backwards_matches (arg_ : string) : M (bool) :=
+ let _s982_ := arg_ in
+ (_s972_ _s982_) >>= fun w__0 : option ((mword 1 * mword 1 * mword 1 * mword 1)) =>
+ (if match w__0 with | Some (i, o, r, w) => true | _ => false end then
+ (_s972_ _s982_) >>= fun w__1 : option ((mword 1 * mword 1 * mword 1 * mword 1)) =>
+ (match w__1 with | Some (i, o, r, w) => returnm true | _ => exit tt : M (bool) end)
: M (bool)
- else returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))})))
+ else returnm false)
: M (bool).
-Definition _s722_ (_s723_ : string)
-: M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string))) :=
-
- (match _s723_ with
- | _s724_ =>
- (bit_maybe_i_matches_prefix _s724_) >>= fun w__0 : option ((mword 1 * {n : Z & ArithFact (n >=
+Definition _s983_ (_s984_ : string) : M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string))) :=
+ (match _s984_ with
+ | _s985_ =>
+ (bit_maybe_i_matches_prefix _s985_) >>= fun w__0 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((i, existT _ _s725_ _)) =>
- (match (string_drop _s724_ _s725_) with
- | _s726_ =>
- (bit_maybe_o_matches_prefix _s726_) >>= fun w__1 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (i, existT _ _s986_ _) =>
+ (match (string_drop _s985_ _s986_) with
+ | _s987_ =>
+ (bit_maybe_o_matches_prefix _s987_) >>= fun w__1 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((o, existT _ _s727_ _)) =>
- (match (string_drop _s726_ _s727_) with
- | _s728_ =>
- (bit_maybe_r_matches_prefix _s728_) >>= fun w__2 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (o, existT _ _s988_ _) =>
+ (match (string_drop _s987_ _s988_) with
+ | _s989_ =>
+ (bit_maybe_r_matches_prefix _s989_) >>= fun w__2 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((r, existT _ _s729_ _)) =>
- (match (string_drop _s728_ _s729_) with
- | _s730_ =>
- (bit_maybe_w_matches_prefix _s730_) >>= fun w__3 : option ((mword 1 * {n : Z & ArithFact (n >=
+ | Some (r, existT _ _s990_ _) =>
+ (match (string_drop _s989_ _s990_) with
+ | _s991_ =>
+ (bit_maybe_w_matches_prefix _s991_) >>= fun w__3 : option ((mword 1 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((w, existT _ _s731_ _)) =>
- match (string_drop _s730_ _s731_) with
- | s_ => Some ((i, o, r, w, s_))
- end
- | _ => None
- end)
- : option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
+ returnm (match w__3 with
+ | Some (w, existT _ _s992_ _) =>
+ match (string_drop _s991_ _s992_) with
+ | s_ => Some (i, o, r, w, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
- | _ =>
- returnm (None
- : option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
- | _ => returnm (None : option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
- | _ => returnm (None : option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
+ | _ => returnm None
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string)))
end)
: M (option ((mword 1 * mword 1 * mword 1 * mword 1 * string))).
-Definition fence_bits_matches_prefix (arg_ : string)
-: M (option ((mword 4 * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s732_ := arg_ in
- (_s722_ _s732_) >>= fun w__0 : option ((mword 1 * mword 1 * mword 1 * mword 1 * string)) =>
- (if ((match w__0 with | Some ((i, o, r, w, s_)) => true | _ => false end)) then
- (_s722_ _s732_) >>= fun w__1 : option ((mword 1 * mword 1 * mword 1 * mword 1 * string)) =>
+Definition fence_bits_matches_prefix (arg_ : string)
+: M (option ((mword 4 * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s993_ := arg_ in
+ (_s983_ _s993_) >>= fun w__0 : option ((mword 1 * mword 1 * mword 1 * mword 1 * string)) =>
+ (if match w__0 with | Some (i, o, r, w, s_) => true | _ => false end then
+ (_s983_ _s993_) >>= fun w__1 : option ((mword 1 * mword 1 * mword 1 * mword 1 * string)) =>
(match w__1 with
- | Some ((i, o, r, w, s_)) =>
- returnm ((Some
- ((concat_vec (i : bits 1)
- (concat_vec (o : bits 1) (concat_vec (r : bits 1) (w : bits 1))), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((mword 4 * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((mword 4 * {n : Z & ArithFact (n >= 0)})))
+ | Some (i, o, r, w, s_) =>
+ returnm (Some
+ (concat_vec (i : bits 1)
+ (concat_vec (o : bits 1) (concat_vec (r : bits 1) (w : bits 1))), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((mword 4 * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((mword 4 * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((mword 4 * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((mword 4 * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((mword 4 * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((mword 4 * {n : Z & ArithFact (n >=? 0)}))).
-Definition aqrl_str (aq : bool) (rl : bool)
-: string :=
-
+Definition aqrl_str (aq : bool) (rl : bool) : string :=
match (aq, rl) with
| (false, false) => ""
| (false, true) => ".rl"
@@ -15936,62 +15058,49 @@ Definition aqrl_str (aq : bool) (rl : bool)
| (true, true) => ".aqrl"
end.
-Definition lrsc_width_str (width : word_width)
-: string :=
-
+Definition lrsc_width_str (width : word_width) : string :=
match width with | BYTE => ".b" | HALF => ".h" | WORD => ".w" | DOUBLE => ".d" end.
Definition process_loadres {n : Z}
(rd : mword 5) (addr : mword 64) (value : MemoryOpResult (mword (8 * n))) (is_unsigned : bool)
-`{ArithFact (0 < n /\ n <= 8)}
+`{ArithFact ((0 <? n) && (n <=? 8))}
: M (Retired) :=
-
(match (extend_value is_unsigned value) with
- | MemValue (result) =>
+ | MemValue result =>
let '_ := (load_reservation addr) : unit in
- (wX (projT1 (regidx_to_regno rd)) result) >> returnm (RETIRE_SUCCESS : Retired)
- | MemException (e) => (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
+ (wX_bits rd result) >> returnm RETIRE_SUCCESS
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
end)
: M (Retired).
-Definition encdec_amoop_forwards (arg_ : amoop)
-: mword 5 :=
-
+Definition encdec_amoop_forwards (arg_ : amoop) : mword 5 :=
match arg_ with
- | AMOSWAP => (vec_of_bits [B0;B0;B0;B0;B1] : mword 5)
- | AMOADD => (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- | AMOXOR => (vec_of_bits [B0;B0;B1;B0;B0] : mword 5)
- | AMOAND => (vec_of_bits [B0;B1;B1;B0;B0] : mword 5)
- | AMOOR => (vec_of_bits [B0;B1;B0;B0;B0] : mword 5)
- | AMOMIN => (vec_of_bits [B1;B0;B0;B0;B0] : mword 5)
- | AMOMAX => (vec_of_bits [B1;B0;B1;B0;B0] : mword 5)
- | AMOMINU => (vec_of_bits [B1;B1;B0;B0;B0] : mword 5)
- | AMOMAXU => (vec_of_bits [B1;B1;B1;B0;B0] : mword 5)
+ | AMOSWAP => 'b"00001" : mword 5
+ | AMOADD => 'b"00000" : mword 5
+ | AMOXOR => 'b"00100" : mword 5
+ | AMOAND => 'b"01100" : mword 5
+ | AMOOR => 'b"01000" : mword 5
+ | AMOMIN => 'b"10000" : mword 5
+ | AMOMAX => 'b"10100" : mword 5
+ | AMOMINU => 'b"11000" : mword 5
+ | AMOMAXU => 'b"11100" : mword 5
end.
-Definition encdec_amoop_backwards (arg_ : mword 5)
-: M (amoop) :=
-
+Definition encdec_amoop_backwards (arg_ : mword 5) : M (amoop) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))) then returnm (AMOSWAP : amoop)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))) then returnm (AMOADD : amoop)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))) then returnm (AMOXOR : amoop)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))) then returnm (AMOAND : amoop)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))) then returnm (AMOOR : amoop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))) then returnm (AMOMIN : amoop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))) then returnm (AMOMAX : amoop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))) then
- returnm (AMOMINU
- : amoop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))) then
- returnm (AMOMAXU
- : amoop)
+ (if eq_vec b__0 ('b"00001" : mword 5) then returnm AMOSWAP
+ else if eq_vec b__0 ('b"00000" : mword 5) then returnm AMOADD
+ else if eq_vec b__0 ('b"00100" : mword 5) then returnm AMOXOR
+ else if eq_vec b__0 ('b"01100" : mword 5) then returnm AMOAND
+ else if eq_vec b__0 ('b"01000" : mword 5) then returnm AMOOR
+ else if eq_vec b__0 ('b"10000" : mword 5) then returnm AMOMIN
+ else if eq_vec b__0 ('b"10100" : mword 5) then returnm AMOMAX
+ else if eq_vec b__0 ('b"11000" : mword 5) then returnm AMOMINU
+ else if eq_vec b__0 ('b"11100" : mword 5) then returnm AMOMAXU
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (amoop).
-Definition encdec_amoop_forwards_matches (arg_ : amoop)
-: bool :=
-
+Definition encdec_amoop_forwards_matches (arg_ : amoop) : bool :=
match arg_ with
| AMOSWAP => true
| AMOADD => true
@@ -16004,24 +15113,20 @@ Definition encdec_amoop_forwards_matches (arg_ : amoop)
| AMOMAXU => true
end.
-Definition encdec_amoop_backwards_matches (arg_ : mword 5)
-: bool :=
-
+Definition encdec_amoop_backwards_matches (arg_ : mword 5) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B1] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B0;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0] : mword 5))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B0;B0] : mword 5))) then true
+ if eq_vec b__0 ('b"00001" : mword 5) then true
+ else if eq_vec b__0 ('b"00000" : mword 5) then true
+ else if eq_vec b__0 ('b"00100" : mword 5) then true
+ else if eq_vec b__0 ('b"01100" : mword 5) then true
+ else if eq_vec b__0 ('b"01000" : mword 5) then true
+ else if eq_vec b__0 ('b"10000" : mword 5) then true
+ else if eq_vec b__0 ('b"10100" : mword 5) then true
+ else if eq_vec b__0 ('b"11000" : mword 5) then true
+ else if eq_vec b__0 ('b"11100" : mword 5) then true
else false.
-Definition amo_mnemonic_forwards (arg_ : amoop)
-: string :=
-
+Definition amo_mnemonic_forwards (arg_ : amoop) : string :=
match arg_ with
| AMOSWAP => "amoswap"
| AMOADD => "amoadd"
@@ -16034,25 +15139,21 @@ Definition amo_mnemonic_forwards (arg_ : amoop)
| AMOMAXU => "amomaxu"
end.
-Definition amo_mnemonic_backwards (arg_ : string)
-: M (amoop) :=
-
+Definition amo_mnemonic_backwards (arg_ : string) : M (amoop) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "amoswap")) then returnm (AMOSWAP : amoop)
- else if ((generic_eq p0_ "amoadd")) then returnm (AMOADD : amoop)
- else if ((generic_eq p0_ "amoxor")) then returnm (AMOXOR : amoop)
- else if ((generic_eq p0_ "amoand")) then returnm (AMOAND : amoop)
- else if ((generic_eq p0_ "amoor")) then returnm (AMOOR : amoop)
- else if ((generic_eq p0_ "amomin")) then returnm (AMOMIN : amoop)
- else if ((generic_eq p0_ "amomax")) then returnm (AMOMAX : amoop)
- else if ((generic_eq p0_ "amominu")) then returnm (AMOMINU : amoop)
- else if ((generic_eq p0_ "amomaxu")) then returnm (AMOMAXU : amoop)
+ (if generic_eq p0_ "amoswap" then returnm AMOSWAP
+ else if generic_eq p0_ "amoadd" then returnm AMOADD
+ else if generic_eq p0_ "amoxor" then returnm AMOXOR
+ else if generic_eq p0_ "amoand" then returnm AMOAND
+ else if generic_eq p0_ "amoor" then returnm AMOOR
+ else if generic_eq p0_ "amomin" then returnm AMOMIN
+ else if generic_eq p0_ "amomax" then returnm AMOMAX
+ else if generic_eq p0_ "amominu" then returnm AMOMINU
+ else if generic_eq p0_ "amomaxu" then returnm AMOMAXU
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (amoop).
-Definition amo_mnemonic_forwards_matches (arg_ : amoop)
-: bool :=
-
+Definition amo_mnemonic_forwards_matches (arg_ : amoop) : bool :=
match arg_ with
| AMOSWAP => true
| AMOADD => true
@@ -16065,225 +15166,189 @@ Definition amo_mnemonic_forwards_matches (arg_ : amoop)
| AMOMAXU => true
end.
-Definition amo_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition amo_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "amoswap")) then true
- else if ((generic_eq p0_ "amoadd")) then true
- else if ((generic_eq p0_ "amoxor")) then true
- else if ((generic_eq p0_ "amoand")) then true
- else if ((generic_eq p0_ "amoor")) then true
- else if ((generic_eq p0_ "amomin")) then true
- else if ((generic_eq p0_ "amomax")) then true
- else if ((generic_eq p0_ "amominu")) then true
- else if ((generic_eq p0_ "amomaxu")) then true
+ if generic_eq p0_ "amoswap" then true
+ else if generic_eq p0_ "amoadd" then true
+ else if generic_eq p0_ "amoxor" then true
+ else if generic_eq p0_ "amoand" then true
+ else if generic_eq p0_ "amoor" then true
+ else if generic_eq p0_ "amomin" then true
+ else if generic_eq p0_ "amomax" then true
+ else if generic_eq p0_ "amominu" then true
+ else if generic_eq p0_ "amomaxu" then true
else false.
-Definition _s765_ (_s766_ : string)
-: option string :=
-
- let _s767_ := _s766_ in
- if ((string_startswith _s767_ "amomaxu")) then
- match (string_drop _s767_ (projT1 (string_length "amomaxu"))) with | s_ => Some (s_) end
+Definition _s1026_ (_s1027_ : string) : option string :=
+ let _s1028_ := _s1027_ in
+ if string_startswith _s1028_ "amomaxu" then
+ match (string_drop _s1028_ (projT1 (string_length "amomaxu"))) with | s_ => Some s_ end
else None.
-Definition _s761_ (_s762_ : string)
-: option string :=
-
- let _s763_ := _s762_ in
- if ((string_startswith _s763_ "amominu")) then
- match (string_drop _s763_ (projT1 (string_length "amominu"))) with | s_ => Some (s_) end
+Definition _s1022_ (_s1023_ : string) : option string :=
+ let _s1024_ := _s1023_ in
+ if string_startswith _s1024_ "amominu" then
+ match (string_drop _s1024_ (projT1 (string_length "amominu"))) with | s_ => Some s_ end
else None.
-Definition _s757_ (_s758_ : string)
-: option string :=
-
- let _s759_ := _s758_ in
- if ((string_startswith _s759_ "amomax")) then
- match (string_drop _s759_ (projT1 (string_length "amomax"))) with | s_ => Some (s_) end
+Definition _s1018_ (_s1019_ : string) : option string :=
+ let _s1020_ := _s1019_ in
+ if string_startswith _s1020_ "amomax" then
+ match (string_drop _s1020_ (projT1 (string_length "amomax"))) with | s_ => Some s_ end
else None.
-Definition _s753_ (_s754_ : string)
-: option string :=
-
- let _s755_ := _s754_ in
- if ((string_startswith _s755_ "amomin")) then
- match (string_drop _s755_ (projT1 (string_length "amomin"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s749_ (_s750_ : string)
-: option string :=
-
- let _s751_ := _s750_ in
- if ((string_startswith _s751_ "amoor")) then
- match (string_drop _s751_ (projT1 (string_length "amoor"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s745_ (_s746_ : string)
-: option string :=
-
- let _s747_ := _s746_ in
- if ((string_startswith _s747_ "amoand")) then
- match (string_drop _s747_ (projT1 (string_length "amoand"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s741_ (_s742_ : string)
-: option string :=
-
- let _s743_ := _s742_ in
- if ((string_startswith _s743_ "amoxor")) then
- match (string_drop _s743_ (projT1 (string_length "amoxor"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s737_ (_s738_ : string)
-: option string :=
-
- let _s739_ := _s738_ in
- if ((string_startswith _s739_ "amoadd")) then
- match (string_drop _s739_ (projT1 (string_length "amoadd"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s733_ (_s734_ : string)
-: option string :=
-
- let _s735_ := _s734_ in
- if ((string_startswith _s735_ "amoswap")) then
- match (string_drop _s735_ (projT1 (string_length "amoswap"))) with | s_ => Some (s_) end
- else None.
-
-Definition amo_mnemonic_matches_prefix (arg_ : string)
-: M (option ((amoop * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s736_ := arg_ in
- (if ((match (_s733_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s733_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOSWAP, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s737_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s737_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOADD, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s741_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s741_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOXOR, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s745_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s745_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOAND, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s749_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s749_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOOR, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s753_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s753_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOMIN, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s757_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s757_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOMAX, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s761_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s761_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOMINU, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s765_ _s736_) with | Some (s_) => true | _ => false end)) then
- (match (_s765_ _s736_) with
- | Some (s_) =>
- returnm ((Some
- ((AMOMAXU, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((amoop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
+Definition _s1014_ (_s1015_ : string) : option string :=
+ let _s1016_ := _s1015_ in
+ if string_startswith _s1016_ "amomin" then
+ match (string_drop _s1016_ (projT1 (string_length "amomin"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1010_ (_s1011_ : string) : option string :=
+ let _s1012_ := _s1011_ in
+ if string_startswith _s1012_ "amoor" then
+ match (string_drop _s1012_ (projT1 (string_length "amoor"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1006_ (_s1007_ : string) : option string :=
+ let _s1008_ := _s1007_ in
+ if string_startswith _s1008_ "amoand" then
+ match (string_drop _s1008_ (projT1 (string_length "amoand"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1002_ (_s1003_ : string) : option string :=
+ let _s1004_ := _s1003_ in
+ if string_startswith _s1004_ "amoxor" then
+ match (string_drop _s1004_ (projT1 (string_length "amoxor"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s998_ (_s999_ : string) : option string :=
+ let _s1000_ := _s999_ in
+ if string_startswith _s1000_ "amoadd" then
+ match (string_drop _s1000_ (projT1 (string_length "amoadd"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s994_ (_s995_ : string) : option string :=
+ let _s996_ := _s995_ in
+ if string_startswith _s996_ "amoswap" then
+ match (string_drop _s996_ (projT1 (string_length "amoswap"))) with | s_ => Some s_ end
+ else None.
+
+Definition amo_mnemonic_matches_prefix (arg_ : string)
+: M (option ((amoop * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s997_ := arg_ in
+ (if match (_s994_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s994_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOSWAP, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s998_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s998_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOADD, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1002_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s1002_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOXOR, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1006_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s1006_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOAND, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1010_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s1010_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOOR, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((amoop * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((amoop * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1014_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s1014_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOMIN, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1018_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s1018_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOMAX, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1022_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s1022_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOMINU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1026_ _s997_) with | Some s_ => true | _ => false end then
+ (match (_s1026_ _s997_) with
+ | Some s_ =>
+ returnm (Some
+ (AMOMAXU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((amoop * {n : Z & ArithFact (n >=? 0)}))).
-Definition encdec_mul_op_forwards (arg_ : (bool * bool * bool))
-: M (mword 3) :=
-
+Definition encdec_mul_op_forwards (arg_ : (bool * bool * bool)) : M (mword 3) :=
(match arg_ with
- | (false, true, true) => returnm ((vec_of_bits [B0;B0;B0] : mword 3) : mword 3)
- | (true, true, true) => returnm ((vec_of_bits [B0;B0;B1] : mword 3) : mword 3)
- | (true, true, false) => returnm ((vec_of_bits [B0;B1;B0] : mword 3) : mword 3)
- | (true, false, false) => returnm ((vec_of_bits [B0;B1;B1] : mword 3) : mword 3)
+ | (false, true, true) => returnm ('b"000" : mword 3)
+ | (true, true, true) => returnm ('b"001" : mword 3)
+ | (true, true, false) => returnm ('b"010" : mword 3)
+ | (true, false, false) => returnm ('b"011" : mword 3)
| _ => exit tt : M (mword 3)
end)
: M (mword 3).
-Definition encdec_mul_op_backwards (arg_ : mword 3)
-: M ((bool * bool * bool)) :=
-
+Definition encdec_mul_op_backwards (arg_ : mword 3) : M ((bool * bool * bool)) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then returnm (false, true, true)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then returnm (true, true, true)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then returnm (true, true, false)
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then returnm (true, false, false)
+ (if eq_vec b__0 ('b"000" : mword 3) then returnm (false, true, true)
+ else if eq_vec b__0 ('b"001" : mword 3) then returnm (true, true, true)
+ else if eq_vec b__0 ('b"010" : mword 3) then returnm (true, true, false)
+ else if eq_vec b__0 ('b"011" : mword 3) then returnm (true, false, false)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M ((bool * bool * bool)).
-Definition encdec_mul_op_forwards_matches (arg_ : (bool * bool * bool))
-: bool :=
-
+Definition encdec_mul_op_forwards_matches (arg_ : (bool * bool * bool)) : bool :=
match arg_ with
| (false, true, true) => true
| (true, true, true) => true
@@ -16292,42 +15357,34 @@ Definition encdec_mul_op_forwards_matches (arg_ : (bool * bool * bool))
| _ => false
end.
-Definition encdec_mul_op_backwards_matches (arg_ : mword 3)
-: bool :=
-
+Definition encdec_mul_op_backwards_matches (arg_ : mword 3) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B0;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B0] : mword 3))) then true
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1] : mword 3))) then true
+ if eq_vec b__0 ('b"000" : mword 3) then true
+ else if eq_vec b__0 ('b"001" : mword 3) then true
+ else if eq_vec b__0 ('b"010" : mword 3) then true
+ else if eq_vec b__0 ('b"011" : mword 3) then true
else false.
-Definition mul_mnemonic_forwards (arg_ : (bool * bool * bool))
-: M (string) :=
-
+Definition mul_mnemonic_forwards (arg_ : (bool * bool * bool)) : M (string) :=
(match arg_ with
- | (false, true, true) => returnm ("mul" : string)
- | (true, true, true) => returnm ("mulh" : string)
- | (true, true, false) => returnm ("mulhsu" : string)
- | (true, false, false) => returnm ("mulhu" : string)
+ | (false, true, true) => returnm "mul"
+ | (true, true, true) => returnm "mulh"
+ | (true, true, false) => returnm "mulhsu"
+ | (true, false, false) => returnm "mulhu"
| _ => exit tt : M (string)
end)
: M (string).
-Definition mul_mnemonic_backwards (arg_ : string)
-: M ((bool * bool * bool)) :=
-
+Definition mul_mnemonic_backwards (arg_ : string) : M ((bool * bool * bool)) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "mul")) then returnm (false, true, true)
- else if ((generic_eq p0_ "mulh")) then returnm (true, true, true)
- else if ((generic_eq p0_ "mulhsu")) then returnm (true, true, false)
- else if ((generic_eq p0_ "mulhu")) then returnm (true, false, false)
+ (if generic_eq p0_ "mul" then returnm (false, true, true)
+ else if generic_eq p0_ "mulh" then returnm (true, true, true)
+ else if generic_eq p0_ "mulhsu" then returnm (true, true, false)
+ else if generic_eq p0_ "mulhu" then returnm (true, false, false)
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M ((bool * bool * bool)).
-Definition mul_mnemonic_forwards_matches (arg_ : (bool * bool * bool))
-: bool :=
-
+Definition mul_mnemonic_forwards_matches (arg_ : (bool * bool * bool)) : bool :=
match arg_ with
| (false, true, true) => true
| (true, true, true) => true
@@ -16336,2780 +15393,7875 @@ Definition mul_mnemonic_forwards_matches (arg_ : (bool * bool * bool))
| _ => false
end.
-Definition mul_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition mul_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "mul")) then true
- else if ((generic_eq p0_ "mulh")) then true
- else if ((generic_eq p0_ "mulhsu")) then true
- else if ((generic_eq p0_ "mulhu")) then true
+ if generic_eq p0_ "mul" then true
+ else if generic_eq p0_ "mulh" then true
+ else if generic_eq p0_ "mulhsu" then true
+ else if generic_eq p0_ "mulhu" then true
else false.
-Definition _s781_ (_s782_ : string)
-: option string :=
-
- let _s783_ := _s782_ in
- if ((string_startswith _s783_ "mulhu")) then
- match (string_drop _s783_ (projT1 (string_length "mulhu"))) with | s_ => Some (s_) end
+Definition _s1042_ (_s1043_ : string) : option string :=
+ let _s1044_ := _s1043_ in
+ if string_startswith _s1044_ "mulhu" then
+ match (string_drop _s1044_ (projT1 (string_length "mulhu"))) with | s_ => Some s_ end
else None.
-Definition _s777_ (_s778_ : string)
-: option string :=
-
- let _s779_ := _s778_ in
- if ((string_startswith _s779_ "mulhsu")) then
- match (string_drop _s779_ (projT1 (string_length "mulhsu"))) with | s_ => Some (s_) end
+Definition _s1038_ (_s1039_ : string) : option string :=
+ let _s1040_ := _s1039_ in
+ if string_startswith _s1040_ "mulhsu" then
+ match (string_drop _s1040_ (projT1 (string_length "mulhsu"))) with | s_ => Some s_ end
else None.
-Definition _s773_ (_s774_ : string)
-: option string :=
-
- let _s775_ := _s774_ in
- if ((string_startswith _s775_ "mulh")) then
- match (string_drop _s775_ (projT1 (string_length "mulh"))) with | s_ => Some (s_) end
+Definition _s1034_ (_s1035_ : string) : option string :=
+ let _s1036_ := _s1035_ in
+ if string_startswith _s1036_ "mulh" then
+ match (string_drop _s1036_ (projT1 (string_length "mulh"))) with | s_ => Some s_ end
else None.
-Definition _s769_ (_s770_ : string)
-: option string :=
-
- let _s771_ := _s770_ in
- if ((string_startswith _s771_ "mul")) then
- match (string_drop _s771_ (projT1 (string_length "mul"))) with | s_ => Some (s_) end
- else None.
-
-Definition mul_mnemonic_matches_prefix (arg_ : string)
-: M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s772_ := arg_ in
- (if ((match (_s769_ _s772_) with | Some (s_) => true | _ => false end)) then
- (match (_s769_ _s772_) with
- | Some (s_) =>
- returnm ((Some
- (((false, true, true), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s773_ _s772_) with | Some (s_) => true | _ => false end)) then
- (match (_s773_ _s772_) with
- | Some (s_) =>
- returnm ((Some
- (((true, true, true), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s777_ _s772_) with | Some (s_) => true | _ => false end)) then
- (match (_s777_ _s772_) with
- | Some (s_) =>
- returnm ((Some
- (((true, true, false), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s781_ _s772_) with | Some (s_) => true | _ => false end)) then
- (match (_s781_ _s772_) with
- | Some (s_) =>
- returnm ((Some
- (((true, false, false), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
+Definition _s1030_ (_s1031_ : string) : option string :=
+ let _s1032_ := _s1031_ in
+ if string_startswith _s1032_ "mul" then
+ match (string_drop _s1032_ (projT1 (string_length "mul"))) with | s_ => Some s_ end
+ else None.
+
+Definition mul_mnemonic_matches_prefix (arg_ : string)
+: M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1033_ := arg_ in
+ (if match (_s1030_ _s1033_) with | Some s_ => true | _ => false end then
+ (match (_s1030_ _s1033_) with
+ | Some s_ =>
+ returnm (Some
+ ((false, true, true), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1034_ _s1033_) with | Some s_ => true | _ => false end then
+ (match (_s1034_ _s1033_) with
+ | Some s_ =>
+ returnm (Some
+ ((true, true, true), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1038_ _s1033_) with | Some s_ => true | _ => false end then
+ (match (_s1038_ _s1033_) with
+ | Some s_ =>
+ returnm (Some
+ ((true, true, false), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1042_ _s1033_) with | Some s_ => true | _ => false end then
+ (match (_s1042_ _s1033_) with
+ | Some s_ =>
+ returnm (Some
+ ((true, false, false), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)}))))
- : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >= 0)}))).
+ : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option (((bool * bool * bool) * {n : Z & ArithFact (n >=? 0)}))).
-Definition maybe_not_u_forwards (arg_ : bool)
-: string :=
-
+Definition maybe_not_u_forwards (arg_ : bool) : string :=
match arg_ with | false => "u" | true => "" end.
-Definition maybe_not_u_backwards (arg_ : string)
-: M (bool) :=
-
+Definition maybe_not_u_backwards (arg_ : string) : M (bool) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "u")) then
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
- else if ((generic_eq p0_ "")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ (if generic_eq p0_ "u" then returnm false
+ else if generic_eq p0_ "" then returnm true
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bool).
-Definition maybe_not_u_forwards_matches (arg_ : bool)
-: bool :=
-
+Definition maybe_not_u_forwards_matches (arg_ : bool) : bool :=
match arg_ with | false => true | true => true end.
-Definition maybe_not_u_backwards_matches (arg_ : string)
-: bool :=
-
+Definition maybe_not_u_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "u")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ "u" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s789_ (_s790_ : string)
-: option string :=
-
- let _s791_ := _s790_ in
- if ((string_startswith _s791_ "")) then
- match (string_drop _s791_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s1050_ (_s1051_ : string) : option string :=
+ let _s1052_ := _s1051_ in
+ if string_startswith _s1052_ "" then
+ match (string_drop _s1052_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s785_ (_s786_ : string)
-: option string :=
-
- let _s787_ := _s786_ in
- if ((string_startswith _s787_ "u")) then
- match (string_drop _s787_ (projT1 (string_length "u"))) with | s_ => Some (s_) end
+Definition _s1046_ (_s1047_ : string) : option string :=
+ let _s1048_ := _s1047_ in
+ if string_startswith _s1048_ "u" then
+ match (string_drop _s1048_ (projT1 (string_length "u"))) with | s_ => Some s_ end
else None.
-Definition maybe_not_u_matches_prefix (arg_ : string)
-: M (option ((bool * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s788_ := arg_ in
- (if ((match (_s785_ _s788_) with | Some (s_) => true | _ => false end)) then
- (match (_s785_ _s788_) with
- | Some (s_) =>
- returnm ((Some
- ((false, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
+Definition maybe_not_u_matches_prefix (arg_ : string)
+: M (option ((bool * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1049_ := arg_ in
+ (if match (_s1046_ _s1049_) with | Some s_ => true | _ => false end then
+ (match (_s1046_ _s1049_) with
+ | Some s_ =>
+ returnm (Some
+ (false, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s789_ _s788_) with | Some (s_) => true | _ => false end)) then
- (match (_s789_ _s788_) with
- | Some (s_) =>
- returnm ((Some
- ((true, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1050_ _s1049_) with | Some s_ => true | _ => false end then
+ (match (_s1050_ _s1049_) with
+ | Some s_ =>
+ returnm (Some
+ (true, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((bool * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((bool * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)}))).
-Definition encdec_csrop_forwards (arg_ : csrop)
-: mword 2 :=
-
+Definition encdec_csrop_forwards (arg_ : csrop) : mword 2 :=
match arg_ with
- | CSRRW => (vec_of_bits [B0;B1] : mword 2)
- | CSRRS => (vec_of_bits [B1;B0] : mword 2)
- | CSRRC => (vec_of_bits [B1;B1] : mword 2)
+ | CSRRW => 'b"01" : mword 2
+ | CSRRS => 'b"10" : mword 2
+ | CSRRC => 'b"11" : mword 2
end.
-Definition encdec_csrop_backwards (arg_ : mword 2)
-: M (csrop) :=
-
+Definition encdec_csrop_backwards (arg_ : mword 2) : M (csrop) :=
let b__0 := arg_ in
- (if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then returnm (CSRRW : csrop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then returnm (CSRRS : csrop)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then returnm (CSRRC : csrop)
+ (if eq_vec b__0 ('b"01" : mword 2) then returnm CSRRW
+ else if eq_vec b__0 ('b"10" : mword 2) then returnm CSRRS
+ else if eq_vec b__0 ('b"11" : mword 2) then returnm CSRRC
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (csrop).
-Definition encdec_csrop_forwards_matches (arg_ : csrop)
-: bool :=
-
+Definition encdec_csrop_forwards_matches (arg_ : csrop) : bool :=
match arg_ with | CSRRW => true | CSRRS => true | CSRRC => true end.
-Definition encdec_csrop_backwards_matches (arg_ : mword 2)
-: bool :=
-
+Definition encdec_csrop_backwards_matches (arg_ : mword 2) : bool :=
let b__0 := arg_ in
- if ((eq_vec b__0 (vec_of_bits [B0;B1] : mword 2))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B0] : mword 2))) then true
- else if ((eq_vec b__0 (vec_of_bits [B1;B1] : mword 2))) then true
+ if eq_vec b__0 ('b"01" : mword 2) then true
+ else if eq_vec b__0 ('b"10" : mword 2) then true
+ else if eq_vec b__0 ('b"11" : mword 2) then true
else false.
-Definition readCSR (csr : mword 12)
-: M (mword 64) :=
-
+Definition readCSR (csr : mword 12) : M (mword 64) :=
(match (csr, 64) with
- | (b__0, g__222) =>
- (if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B0;B1] : mword 12))) then
+ | (b__0, g__324) =>
+ (if eq_vec b__0 (Ox"F11" : mword 12) then
((read_reg mvendorid_ref) : M (mword 32)) >>= fun w__0 : mword 32 =>
- returnm ((EXTZ 64 w__0)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B0] : mword 12))) then
+ returnm (EXTZ 64 w__0)
+ else if eq_vec b__0 (Ox"F12" : mword 12) then
((read_reg marchid_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B0;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"F13" : mword 12) then
((read_reg mimpid_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B1;B1;B0;B0;B0;B1;B0;B1;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"F14" : mword 12) then
((read_reg mhartid_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- read_reg mstatus_ref >>= fun w__4 : Mstatus =>
- returnm ((_get_Mstatus_bits w__4)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- read_reg misa_ref >>= fun w__5 : Misa => returnm ((_get_Misa_bits w__5) : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg medeleg_ref >>= fun w__6 : Medeleg =>
- returnm ((_get_Medeleg_bits w__6)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- read_reg mideleg_ref >>= fun w__7 : Minterrupts =>
- returnm ((_get_Minterrupts_bits w__7)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- read_reg mie_ref >>= fun w__8 : Minterrupts =>
- returnm ((_get_Minterrupts_bits w__8)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- (get_mtvec tt)
- : M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"300" : mword 12) then
+ read_reg mstatus_ref >>= fun w__4 : Mstatus => returnm (_get_Mstatus_bits w__4)
+ else if eq_vec b__0 (Ox"301" : mword 12) then
+ read_reg misa_ref >>= fun w__5 : Misa => returnm (_get_Misa_bits w__5)
+ else if eq_vec b__0 (Ox"302" : mword 12) then
+ read_reg medeleg_ref >>= fun w__6 : Medeleg => returnm (_get_Medeleg_bits w__6)
+ else if eq_vec b__0 (Ox"303" : mword 12) then
+ read_reg mideleg_ref >>= fun w__7 : Minterrupts => returnm (_get_Minterrupts_bits w__7)
+ else if eq_vec b__0 (Ox"304" : mword 12) then
+ read_reg mie_ref >>= fun w__8 : Minterrupts => returnm (_get_Minterrupts_bits w__8)
+ else if eq_vec b__0 (Ox"305" : mword 12) then (get_mtvec tt) : M (mword 64)
+ else if eq_vec b__0 (Ox"306" : mword 12) then
read_reg mcounteren_ref >>= fun w__10 : Counteren =>
- returnm ((EXTZ 64 (_get_Counteren_bits w__10))
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ returnm (EXTZ 64 (_get_Counteren_bits w__10))
+ else if eq_vec b__0 (Ox"320" : mword 12) then
+ read_reg mcountinhibit_ref >>= fun w__11 : Counterin =>
+ returnm (EXTZ 64 (_get_Counterin_bits w__11))
+ else if eq_vec b__0 (Ox"340" : mword 12) then
((read_reg mscratch_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- (get_xret_target Machine) >>= fun w__12 : mword 64 =>
- (pc_alignment_mask tt) >>= fun w__13 : mword 64 =>
- returnm ((and_vec w__12 w__13)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg mcause_ref >>= fun w__14 : Mcause =>
- returnm ((_get_Mcause_bits w__14)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"341" : mword 12) then
+ (get_xret_target Machine) >>= fun w__13 : mword 64 =>
+ (pc_alignment_mask tt) >>= fun w__14 : mword 64 => returnm (and_vec w__13 w__14)
+ else if eq_vec b__0 (Ox"342" : mword 12) then
+ read_reg mcause_ref >>= fun w__15 : Mcause => returnm (_get_Mcause_bits w__15)
+ else if eq_vec b__0 (Ox"343" : mword 12) then
((read_reg mtval_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- read_reg mip_ref >>= fun w__16 : Minterrupts =>
- returnm ((_get_Minterrupts_bits w__16)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- (pmpReadCfgReg 0)
- : M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
- (pmpReadCfgReg 2)
- : M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"344" : mword 12) then
+ read_reg mip_ref >>= fun w__17 : Minterrupts => returnm (_get_Minterrupts_bits w__17)
+ else if eq_vec b__0 (Ox"3A0" : mword 12) then (pmpReadCfgReg 0) : M (mword 64)
+ else if eq_vec b__0 (Ox"3A2" : mword 12) then (pmpReadCfgReg 2) : M (mword 64)
+ else if eq_vec b__0 (Ox"3B0" : mword 12) then
((read_reg pmpaddr0_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B1" : mword 12) then
((read_reg pmpaddr1_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B2" : mword 12) then
((read_reg pmpaddr2_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B3" : mword 12) then
((read_reg pmpaddr3_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B4" : mword 12) then
((read_reg pmpaddr4_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B5" : mword 12) then
((read_reg pmpaddr5_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B6" : mword 12) then
((read_reg pmpaddr6_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B7" : mword 12) then
((read_reg pmpaddr7_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B8" : mword 12) then
((read_reg pmpaddr8_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"3B9" : mword 12) then
((read_reg pmpaddr9_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"3BA" : mword 12) then
((read_reg pmpaddr10_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"3BB" : mword 12) then
((read_reg pmpaddr11_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"3BC" : mword 12) then
((read_reg pmpaddr12_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"3BD" : mword 12) then
((read_reg pmpaddr13_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"3BE" : mword 12) then
((read_reg pmpaddr14_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"3BF" : mword 12) then
((read_reg pmpaddr15_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- ((read_reg mcycle_ref) : M (mword 64)) >>= fun w__35 : mword 64 =>
- returnm ((subrange_vec_dec w__35 (Z.sub 64 1) 0)
- : mword (63 - 0 + 1))
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- ((read_reg minstret_ref) : M (mword 64)) >>= fun w__36 : mword 64 =>
- returnm ((subrange_vec_dec w__36 (Z.sub 64 1) 0)
- : mword (63 - 0 + 1))
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- ((read_reg tselect_ref) : M (mword 64)) >>= fun w__37 : mword 64 =>
- returnm ((not_vec w__37)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- read_reg mstatus_ref >>= fun w__38 : Mstatus =>
- returnm ((_get_Sstatus_bits (lower_mstatus w__38))
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg sedeleg_ref >>= fun w__39 : Sedeleg =>
- returnm ((_get_Sedeleg_bits w__39)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
- read_reg sideleg_ref >>= fun w__40 : Sinterrupts =>
- returnm ((_get_Sinterrupts_bits w__40)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- read_reg mie_ref >>= fun w__41 : Minterrupts =>
- read_reg mideleg_ref >>= fun w__42 : Minterrupts =>
- returnm ((_get_Sinterrupts_bits (lower_mie w__41 w__42))
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- (get_stvec tt)
- : M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
- read_reg scounteren_ref >>= fun w__44 : Counteren =>
- returnm ((EXTZ 64 (_get_Counteren_bits w__44))
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"B00" : mword 12) then
+ ((read_reg mcycle_ref) : M (mword 64)) >>= fun w__36 : mword 64 =>
+ returnm (subrange_vec_dec w__36 (Z.sub 64 1) 0)
+ else if eq_vec b__0 (Ox"B02" : mword 12) then
+ ((read_reg minstret_ref) : M (mword 64)) >>= fun w__37 : mword 64 =>
+ returnm (subrange_vec_dec w__37 (Z.sub 64 1) 0)
+ else if eq_vec b__0 (Ox"7A0" : mword 12) then
+ ((read_reg tselect_ref) : M (mword 64)) >>= fun w__38 : mword 64 =>
+ returnm (not_vec w__38)
+ else if eq_vec b__0 (Ox"100" : mword 12) then
+ read_reg mstatus_ref >>= fun w__39 : Mstatus =>
+ returnm (_get_Sstatus_bits (lower_mstatus w__39))
+ else if eq_vec b__0 (Ox"102" : mword 12) then
+ read_reg sedeleg_ref >>= fun w__40 : Sedeleg => returnm (_get_Sedeleg_bits w__40)
+ else if eq_vec b__0 (Ox"103" : mword 12) then
+ read_reg sideleg_ref >>= fun w__41 : Sinterrupts => returnm (_get_Sinterrupts_bits w__41)
+ else if eq_vec b__0 (Ox"104" : mword 12) then
+ read_reg mie_ref >>= fun w__42 : Minterrupts =>
+ read_reg mideleg_ref >>= fun w__43 : Minterrupts =>
+ returnm (_get_Sinterrupts_bits (lower_mie w__42 w__43))
+ else if eq_vec b__0 (Ox"105" : mword 12) then (get_stvec tt) : M (mword 64)
+ else if eq_vec b__0 (Ox"106" : mword 12) then
+ read_reg scounteren_ref >>= fun w__45 : Counteren =>
+ returnm (EXTZ 64 (_get_Counteren_bits w__45))
+ else if eq_vec b__0 (Ox"140" : mword 12) then
((read_reg sscratch_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- (get_xret_target Supervisor) >>= fun w__46 : mword 64 =>
- (pc_alignment_mask tt) >>= fun w__47 : mword 64 =>
- returnm ((and_vec w__46 w__47)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg scause_ref >>= fun w__48 : Mcause =>
- returnm ((_get_Mcause_bits w__48)
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ else if eq_vec b__0 (Ox"141" : mword 12) then
+ (get_xret_target Supervisor) >>= fun w__47 : mword 64 =>
+ (pc_alignment_mask tt) >>= fun w__48 : mword 64 => returnm (and_vec w__47 w__48)
+ else if eq_vec b__0 (Ox"142" : mword 12) then
+ read_reg scause_ref >>= fun w__49 : Mcause => returnm (_get_Mcause_bits w__49)
+ else if eq_vec b__0 (Ox"143" : mword 12) then
((read_reg stval_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- read_reg mip_ref >>= fun w__50 : Minterrupts =>
- read_reg mideleg_ref >>= fun w__51 : Minterrupts =>
- returnm ((_get_Sinterrupts_bits (lower_mip w__50 w__51))
- : mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ else if eq_vec b__0 (Ox"144" : mword 12) then
+ read_reg mip_ref >>= fun w__51 : Minterrupts =>
+ read_reg mideleg_ref >>= fun w__52 : Minterrupts =>
+ returnm (_get_Sinterrupts_bits (lower_mip w__51 w__52))
+ else if eq_vec b__0 (Ox"180" : mword 12) then
((read_reg satp_ref) : M (mword 64))
: M (mword 64)
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- ((read_reg mcycle_ref) : M (mword 64)) >>= fun w__53 : mword 64 =>
- returnm ((subrange_vec_dec w__53 (Z.sub 64 1) 0)
- : mword (63 - 0 + 1))
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
- ((read_reg mtime_ref) : M (mword 64)) >>= fun w__54 : mword 64 =>
- returnm ((subrange_vec_dec w__54 (Z.sub 64 1) 0)
- : mword (63 - 0 + 1))
- else if ((eq_vec b__0 (vec_of_bits [B1;B1;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- ((read_reg minstret_ref) : M (mword 64)) >>= fun w__55 : mword 64 =>
- returnm ((subrange_vec_dec w__55 (Z.sub 64 1) 0)
- : mword (63 - 0 + 1))
+ else if eq_vec b__0 (Ox"C00" : mword 12) then
+ ((read_reg mcycle_ref) : M (mword 64)) >>= fun w__54 : mword 64 =>
+ returnm (subrange_vec_dec w__54 (Z.sub 64 1) 0)
+ else if eq_vec b__0 (Ox"C01" : mword 12) then
+ ((read_reg mtime_ref) : M (mword 64)) >>= fun w__55 : mword 64 =>
+ returnm (subrange_vec_dec w__55 (Z.sub 64 1) 0)
+ else if eq_vec b__0 (Ox"C02" : mword 12) then
+ ((read_reg minstret_ref) : M (mword 64)) >>= fun w__56 : mword 64 =>
+ returnm (subrange_vec_dec w__56 (Z.sub 64 1) 0)
else
- (ext_read_CSR csr) >>= fun w__56 : option (mword 64) =>
- returnm ((match w__56 with
- | Some (res) => res
- | None =>
- let '_ := (print_bits "unhandled read to CSR " csr) : unit in
- EXTZ 64 (vec_of_bits [B0;B0;B0;B0] : mword 4)
- end)
- : mword 64))
+ (ext_read_CSR csr) >>= fun w__57 : option (mword 64) =>
+ returnm (match w__57 with
+ | Some res => res
+ | None =>
+ let '_ := (print_bits "unhandled read to CSR " csr) : unit in
+ EXTZ 64 (Ox"0" : mword 4)
+ end))
: M (mword 64)
end) >>= fun res : xlenbits =>
let '_ :=
- (if ((get_config_print_reg tt)) then
+ (if get_config_print_reg tt then
print_endline
(String.append "CSR "
(String.append (csr_name csr) (String.append " -> " (string_of_bits res))))
else tt)
: unit in
- returnm (res
- : mword 64).
+ returnm res.
-Definition writeCSR (csr : mword 12) (value : mword 64)
-: M (unit) :=
-
+Definition writeCSR (csr : mword 12) (value : mword 64) : M (unit) :=
(match (csr, 64) with
- | (b__0, g__221) =>
- (if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ | (b__0, g__323) =>
+ (if eq_vec b__0 (Ox"300" : mword 12) then
read_reg mstatus_ref >>= fun w__0 : Mstatus =>
(legalize_mstatus w__0 value) >>= fun w__1 : Mstatus =>
write_reg mstatus_ref w__1 >>
- read_reg mstatus_ref >>= fun w__2 : Mstatus =>
- returnm ((Some
- (_get_Mstatus_bits w__2))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12))) then
+ read_reg mstatus_ref >>= fun w__2 : Mstatus => returnm (Some (_get_Mstatus_bits w__2))
+ else if eq_vec b__0 (Ox"301" : mword 12) then
read_reg misa_ref >>= fun w__3 : Misa =>
(legalize_misa w__3 value) >>= fun w__4 : Misa =>
write_reg misa_ref w__4 >>
- read_reg misa_ref >>= fun w__5 : Misa =>
- returnm ((Some
- (_get_Misa_bits w__5))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ read_reg misa_ref >>= fun w__5 : Misa => returnm (Some (_get_Misa_bits w__5))
+ else if eq_vec b__0 (Ox"302" : mword 12) then
read_reg medeleg_ref >>= fun w__6 : Medeleg =>
write_reg medeleg_ref (legalize_medeleg w__6 value) >>
- read_reg medeleg_ref >>= fun w__7 : Medeleg =>
- returnm ((Some
- (_get_Medeleg_bits w__7))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ read_reg medeleg_ref >>= fun w__7 : Medeleg => returnm (Some (_get_Medeleg_bits w__7))
+ else if eq_vec b__0 (Ox"303" : mword 12) then
read_reg mideleg_ref >>= fun w__8 : Minterrupts =>
write_reg mideleg_ref (legalize_mideleg w__8 value) >>
read_reg mideleg_ref >>= fun w__9 : Minterrupts =>
- returnm ((Some
- (_get_Minterrupts_bits w__9))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
+ returnm (Some (_get_Minterrupts_bits w__9))
+ else if eq_vec b__0 (Ox"304" : mword 12) then
read_reg mie_ref >>= fun w__10 : Minterrupts =>
(legalize_mie w__10 value) >>= fun w__11 : Minterrupts =>
write_reg mie_ref w__11 >>
read_reg mie_ref >>= fun w__12 : Minterrupts =>
- returnm ((Some
- (_get_Minterrupts_bits w__12))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- (set_mtvec value) >>= fun w__13 : mword 64 =>
- returnm ((Some
- (w__13))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
+ returnm (Some (_get_Minterrupts_bits w__12))
+ else if eq_vec b__0 (Ox"305" : mword 12) then
+ (set_mtvec value) >>= fun w__13 : mword 64 => returnm (Some w__13)
+ else if eq_vec b__0 (Ox"306" : mword 12) then
read_reg mcounteren_ref >>= fun w__14 : Counteren =>
- (legalize_mcounteren w__14 value) >>= fun w__15 : Counteren =>
- write_reg mcounteren_ref w__15 >>
- read_reg mcounteren_ref >>= fun w__16 : Counteren =>
- returnm ((Some
- (EXTZ 64 (_get_Counteren_bits w__16)))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ write_reg mcounteren_ref (legalize_mcounteren w__14 value) >>
+ read_reg mcounteren_ref >>= fun w__15 : Counteren =>
+ returnm (Some (EXTZ 64 (_get_Counteren_bits w__15)))
+ else if eq_vec b__0 (Ox"320" : mword 12) then
+ read_reg mcountinhibit_ref >>= fun w__16 : Counterin =>
+ write_reg mcountinhibit_ref (legalize_mcountinhibit w__16 value) >>
+ read_reg mcountinhibit_ref >>= fun w__17 : Counterin =>
+ returnm (Some (EXTZ 64 (_get_Counterin_bits w__17)))
+ else if eq_vec b__0 (Ox"340" : mword 12) then
write_reg mscratch_ref value >>
- ((read_reg mscratch_ref) : M (mword 64)) >>= fun w__17 : mword 64 =>
- returnm ((Some
- (w__17))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- (set_xret_target Machine value) >>= fun w__18 : mword 64 =>
- returnm ((Some
- (w__18))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ ((read_reg mscratch_ref) : M (mword 64)) >>= fun w__18 : mword 64 => returnm (Some w__18)
+ else if eq_vec b__0 (Ox"341" : mword 12) then
+ (set_xret_target Machine value) >>= fun w__19 : mword 64 => returnm (Some w__19)
+ else if eq_vec b__0 (Ox"342" : mword 12) then
(_set_Mcause_bits mcause_ref value) >>
- read_reg mcause_ref >>= fun w__19 : Mcause =>
- returnm ((Some
- (_get_Mcause_bits w__19))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ read_reg mcause_ref >>= fun w__20 : Mcause => returnm (Some (_get_Mcause_bits w__20))
+ else if eq_vec b__0 (Ox"343" : mword 12) then
write_reg mtval_ref value >>
- ((read_reg mtval_ref) : M (mword 64)) >>= fun w__20 : mword 64 =>
- returnm ((Some
- (w__20))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- read_reg mip_ref >>= fun w__21 : Minterrupts =>
- (legalize_mip w__21 value) >>= fun w__22 : Minterrupts =>
- write_reg mip_ref w__22 >>
- read_reg mip_ref >>= fun w__23 : Minterrupts =>
- returnm ((Some
- (_get_Minterrupts_bits w__23))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
- (pmpWriteCfgReg 0 value) >> returnm ((Some (value)) : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B0;B0;B0;B1;B0] : mword 12))) then
- (pmpWriteCfgReg 2 value) >> returnm ((Some (value)) : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B0] : mword 12))) then
- read_reg pmp0cfg_ref >>= fun w__24 : Pmpcfg_ent =>
- ((read_reg pmpaddr0_ref) : M (mword 64)) >>= fun w__25 : mword 64 =>
- write_reg pmpaddr0_ref (pmpWriteAddr w__24 w__25 value) >>
- ((read_reg pmpaddr0_ref) : M (mword 64)) >>= fun w__26 : mword 64 =>
- returnm ((Some
- (w__26))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B0;B1] : mword 12))) then
- read_reg pmp1cfg_ref >>= fun w__27 : Pmpcfg_ent =>
- ((read_reg pmpaddr1_ref) : M (mword 64)) >>= fun w__28 : mword 64 =>
- write_reg pmpaddr1_ref (pmpWriteAddr w__27 w__28 value) >>
- ((read_reg pmpaddr1_ref) : M (mword 64)) >>= fun w__29 : mword 64 =>
- returnm ((Some
- (w__29))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B0] : mword 12))) then
- read_reg pmp2cfg_ref >>= fun w__30 : Pmpcfg_ent =>
- ((read_reg pmpaddr2_ref) : M (mword 64)) >>= fun w__31 : mword 64 =>
- write_reg pmpaddr2_ref (pmpWriteAddr w__30 w__31 value) >>
- ((read_reg pmpaddr2_ref) : M (mword 64)) >>= fun w__32 : mword 64 =>
- returnm ((Some
- (w__32))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B0;B1;B1] : mword 12))) then
- read_reg pmp3cfg_ref >>= fun w__33 : Pmpcfg_ent =>
- ((read_reg pmpaddr3_ref) : M (mword 64)) >>= fun w__34 : mword 64 =>
- write_reg pmpaddr3_ref (pmpWriteAddr w__33 w__34 value) >>
- ((read_reg pmpaddr3_ref) : M (mword 64)) >>= fun w__35 : mword 64 =>
- returnm ((Some
- (w__35))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B0] : mword 12))) then
- read_reg pmp4cfg_ref >>= fun w__36 : Pmpcfg_ent =>
- ((read_reg pmpaddr4_ref) : M (mword 64)) >>= fun w__37 : mword 64 =>
- write_reg pmpaddr4_ref (pmpWriteAddr w__36 w__37 value) >>
- ((read_reg pmpaddr4_ref) : M (mword 64)) >>= fun w__38 : mword 64 =>
- returnm ((Some
- (w__38))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B0;B1] : mword 12))) then
- read_reg pmp5cfg_ref >>= fun w__39 : Pmpcfg_ent =>
- ((read_reg pmpaddr5_ref) : M (mword 64)) >>= fun w__40 : mword 64 =>
- write_reg pmpaddr5_ref (pmpWriteAddr w__39 w__40 value) >>
- ((read_reg pmpaddr5_ref) : M (mword 64)) >>= fun w__41 : mword 64 =>
- returnm ((Some
- (w__41))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B0] : mword 12))) then
- read_reg pmp6cfg_ref >>= fun w__42 : Pmpcfg_ent =>
- ((read_reg pmpaddr6_ref) : M (mword 64)) >>= fun w__43 : mword 64 =>
- write_reg pmpaddr6_ref (pmpWriteAddr w__42 w__43 value) >>
- ((read_reg pmpaddr6_ref) : M (mword 64)) >>= fun w__44 : mword 64 =>
- returnm ((Some
- (w__44))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B0;B1;B1;B1] : mword 12))) then
- read_reg pmp7cfg_ref >>= fun w__45 : Pmpcfg_ent =>
- ((read_reg pmpaddr7_ref) : M (mword 64)) >>= fun w__46 : mword 64 =>
- write_reg pmpaddr7_ref (pmpWriteAddr w__45 w__46 value) >>
- ((read_reg pmpaddr7_ref) : M (mword 64)) >>= fun w__47 : mword 64 =>
- returnm ((Some
- (w__47))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B0] : mword 12))) then
- read_reg pmp8cfg_ref >>= fun w__48 : Pmpcfg_ent =>
- ((read_reg pmpaddr8_ref) : M (mword 64)) >>= fun w__49 : mword 64 =>
- write_reg pmpaddr8_ref (pmpWriteAddr w__48 w__49 value) >>
- ((read_reg pmpaddr8_ref) : M (mword 64)) >>= fun w__50 : mword 64 =>
- returnm ((Some
- (w__50))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B0;B1] : mword 12))) then
- read_reg pmp9cfg_ref >>= fun w__51 : Pmpcfg_ent =>
- ((read_reg pmpaddr9_ref) : M (mword 64)) >>= fun w__52 : mword 64 =>
- write_reg pmpaddr9_ref (pmpWriteAddr w__51 w__52 value) >>
- ((read_reg pmpaddr9_ref) : M (mword 64)) >>= fun w__53 : mword 64 =>
- returnm ((Some
- (w__53))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B0] : mword 12))) then
- read_reg pmp10cfg_ref >>= fun w__54 : Pmpcfg_ent =>
- ((read_reg pmpaddr10_ref) : M (mword 64)) >>= fun w__55 : mword 64 =>
- write_reg pmpaddr10_ref (pmpWriteAddr w__54 w__55 value) >>
- ((read_reg pmpaddr10_ref) : M (mword 64)) >>= fun w__56 : mword 64 =>
- returnm ((Some
- (w__56))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B0;B1;B1] : mword 12))) then
- read_reg pmp11cfg_ref >>= fun w__57 : Pmpcfg_ent =>
- ((read_reg pmpaddr11_ref) : M (mword 64)) >>= fun w__58 : mword 64 =>
- write_reg pmpaddr11_ref (pmpWriteAddr w__57 w__58 value) >>
- ((read_reg pmpaddr11_ref) : M (mword 64)) >>= fun w__59 : mword 64 =>
- returnm ((Some
- (w__59))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B0] : mword 12))) then
- read_reg pmp12cfg_ref >>= fun w__60 : Pmpcfg_ent =>
- ((read_reg pmpaddr12_ref) : M (mword 64)) >>= fun w__61 : mword 64 =>
- write_reg pmpaddr12_ref (pmpWriteAddr w__60 w__61 value) >>
- ((read_reg pmpaddr12_ref) : M (mword 64)) >>= fun w__62 : mword 64 =>
- returnm ((Some
- (w__62))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B0;B1] : mword 12))) then
- read_reg pmp13cfg_ref >>= fun w__63 : Pmpcfg_ent =>
- ((read_reg pmpaddr13_ref) : M (mword 64)) >>= fun w__64 : mword 64 =>
- write_reg pmpaddr13_ref (pmpWriteAddr w__63 w__64 value) >>
- ((read_reg pmpaddr13_ref) : M (mword 64)) >>= fun w__65 : mword 64 =>
- returnm ((Some
- (w__65))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B0] : mword 12))) then
- read_reg pmp14cfg_ref >>= fun w__66 : Pmpcfg_ent =>
- ((read_reg pmpaddr14_ref) : M (mword 64)) >>= fun w__67 : mword 64 =>
- write_reg pmpaddr14_ref (pmpWriteAddr w__66 w__67 value) >>
- ((read_reg pmpaddr14_ref) : M (mword 64)) >>= fun w__68 : mword 64 =>
- returnm ((Some
- (w__68))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B1;B1;B1;B0;B1;B1;B1;B1;B1;B1] : mword 12))) then
- read_reg pmp15cfg_ref >>= fun w__69 : Pmpcfg_ent =>
- ((read_reg pmpaddr15_ref) : M (mword 64)) >>= fun w__70 : mword 64 =>
- write_reg pmpaddr15_ref (pmpWriteAddr w__69 w__70 value) >>
- ((read_reg pmpaddr15_ref) : M (mword 64)) >>= fun w__71 : mword 64 =>
- returnm ((Some
- (w__71))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- ((read_reg mcycle_ref) : M (mword 64)) >>= fun w__72 : mword 64 =>
- write_reg mcycle_ref (update_subrange_vec_dec w__72 (Z.sub 64 1) 0 value) >>
- returnm ((Some
- (value))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B1;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- ((read_reg minstret_ref) : M (mword 64)) >>= fun w__73 : mword 64 =>
- write_reg minstret_ref (update_subrange_vec_dec w__73 (Z.sub 64 1) 0 value) >>
- write_reg minstret_written_ref true >> returnm ((Some (value)) : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B1;B1;B1;B1;B0;B1;B0;B0;B0;B0;B0] : mword 12))) then
+ ((read_reg mtval_ref) : M (mword 64)) >>= fun w__21 : mword 64 => returnm (Some w__21)
+ else if eq_vec b__0 (Ox"344" : mword 12) then
+ read_reg mip_ref >>= fun w__22 : Minterrupts =>
+ (legalize_mip w__22 value) >>= fun w__23 : Minterrupts =>
+ write_reg mip_ref w__23 >>
+ read_reg mip_ref >>= fun w__24 : Minterrupts =>
+ returnm (Some (_get_Minterrupts_bits w__24))
+ else if eq_vec b__0 (Ox"3A0" : mword 12) then
+ (pmpWriteCfgReg 0 value) >>
+ (pmpReadCfgReg 0) >>= fun w__25 : mword 64 => returnm (Some w__25)
+ else if eq_vec b__0 (Ox"3A2" : mword 12) then
+ (pmpWriteCfgReg 2 value) >>
+ (pmpReadCfgReg 2) >>= fun w__26 : mword 64 => returnm (Some w__26)
+ else if eq_vec b__0 (Ox"3B0" : mword 12) then
+ read_reg pmp0cfg_ref >>= fun w__27 : Pmpcfg_ent =>
+ read_reg pmp1cfg_ref >>= fun w__28 : Pmpcfg_ent =>
+ (pmpTORLocked w__28) >>= fun w__29 : bool =>
+ ((read_reg pmpaddr0_ref) : M (mword 64)) >>= fun w__30 : mword 64 =>
+ write_reg pmpaddr0_ref (pmpWriteAddr (pmpLocked w__27) w__29 w__30 value) >>
+ ((read_reg pmpaddr0_ref) : M (mword 64)) >>= fun w__31 : mword 64 => returnm (Some w__31)
+ else if eq_vec b__0 (Ox"3B1" : mword 12) then
+ read_reg pmp1cfg_ref >>= fun w__32 : Pmpcfg_ent =>
+ read_reg pmp2cfg_ref >>= fun w__33 : Pmpcfg_ent =>
+ (pmpTORLocked w__33) >>= fun w__34 : bool =>
+ ((read_reg pmpaddr1_ref) : M (mword 64)) >>= fun w__35 : mword 64 =>
+ write_reg pmpaddr1_ref (pmpWriteAddr (pmpLocked w__32) w__34 w__35 value) >>
+ ((read_reg pmpaddr1_ref) : M (mword 64)) >>= fun w__36 : mword 64 => returnm (Some w__36)
+ else if eq_vec b__0 (Ox"3B2" : mword 12) then
+ read_reg pmp2cfg_ref >>= fun w__37 : Pmpcfg_ent =>
+ read_reg pmp3cfg_ref >>= fun w__38 : Pmpcfg_ent =>
+ (pmpTORLocked w__38) >>= fun w__39 : bool =>
+ ((read_reg pmpaddr2_ref) : M (mword 64)) >>= fun w__40 : mword 64 =>
+ write_reg pmpaddr2_ref (pmpWriteAddr (pmpLocked w__37) w__39 w__40 value) >>
+ ((read_reg pmpaddr2_ref) : M (mword 64)) >>= fun w__41 : mword 64 => returnm (Some w__41)
+ else if eq_vec b__0 (Ox"3B3" : mword 12) then
+ read_reg pmp3cfg_ref >>= fun w__42 : Pmpcfg_ent =>
+ read_reg pmp4cfg_ref >>= fun w__43 : Pmpcfg_ent =>
+ (pmpTORLocked w__43) >>= fun w__44 : bool =>
+ ((read_reg pmpaddr3_ref) : M (mword 64)) >>= fun w__45 : mword 64 =>
+ write_reg pmpaddr3_ref (pmpWriteAddr (pmpLocked w__42) w__44 w__45 value) >>
+ ((read_reg pmpaddr3_ref) : M (mword 64)) >>= fun w__46 : mword 64 => returnm (Some w__46)
+ else if eq_vec b__0 (Ox"3B4" : mword 12) then
+ read_reg pmp4cfg_ref >>= fun w__47 : Pmpcfg_ent =>
+ read_reg pmp5cfg_ref >>= fun w__48 : Pmpcfg_ent =>
+ (pmpTORLocked w__48) >>= fun w__49 : bool =>
+ ((read_reg pmpaddr4_ref) : M (mword 64)) >>= fun w__50 : mword 64 =>
+ write_reg pmpaddr4_ref (pmpWriteAddr (pmpLocked w__47) w__49 w__50 value) >>
+ ((read_reg pmpaddr4_ref) : M (mword 64)) >>= fun w__51 : mword 64 => returnm (Some w__51)
+ else if eq_vec b__0 (Ox"3B5" : mword 12) then
+ read_reg pmp5cfg_ref >>= fun w__52 : Pmpcfg_ent =>
+ read_reg pmp6cfg_ref >>= fun w__53 : Pmpcfg_ent =>
+ (pmpTORLocked w__53) >>= fun w__54 : bool =>
+ ((read_reg pmpaddr5_ref) : M (mword 64)) >>= fun w__55 : mword 64 =>
+ write_reg pmpaddr5_ref (pmpWriteAddr (pmpLocked w__52) w__54 w__55 value) >>
+ ((read_reg pmpaddr5_ref) : M (mword 64)) >>= fun w__56 : mword 64 => returnm (Some w__56)
+ else if eq_vec b__0 (Ox"3B6" : mword 12) then
+ read_reg pmp6cfg_ref >>= fun w__57 : Pmpcfg_ent =>
+ read_reg pmp7cfg_ref >>= fun w__58 : Pmpcfg_ent =>
+ (pmpTORLocked w__58) >>= fun w__59 : bool =>
+ ((read_reg pmpaddr6_ref) : M (mword 64)) >>= fun w__60 : mword 64 =>
+ write_reg pmpaddr6_ref (pmpWriteAddr (pmpLocked w__57) w__59 w__60 value) >>
+ ((read_reg pmpaddr6_ref) : M (mword 64)) >>= fun w__61 : mword 64 => returnm (Some w__61)
+ else if eq_vec b__0 (Ox"3B7" : mword 12) then
+ read_reg pmp7cfg_ref >>= fun w__62 : Pmpcfg_ent =>
+ read_reg pmp8cfg_ref >>= fun w__63 : Pmpcfg_ent =>
+ (pmpTORLocked w__63) >>= fun w__64 : bool =>
+ ((read_reg pmpaddr7_ref) : M (mword 64)) >>= fun w__65 : mword 64 =>
+ write_reg pmpaddr7_ref (pmpWriteAddr (pmpLocked w__62) w__64 w__65 value) >>
+ ((read_reg pmpaddr7_ref) : M (mword 64)) >>= fun w__66 : mword 64 => returnm (Some w__66)
+ else if eq_vec b__0 (Ox"3B8" : mword 12) then
+ read_reg pmp8cfg_ref >>= fun w__67 : Pmpcfg_ent =>
+ read_reg pmp9cfg_ref >>= fun w__68 : Pmpcfg_ent =>
+ (pmpTORLocked w__68) >>= fun w__69 : bool =>
+ ((read_reg pmpaddr8_ref) : M (mword 64)) >>= fun w__70 : mword 64 =>
+ write_reg pmpaddr8_ref (pmpWriteAddr (pmpLocked w__67) w__69 w__70 value) >>
+ ((read_reg pmpaddr8_ref) : M (mword 64)) >>= fun w__71 : mword 64 => returnm (Some w__71)
+ else if eq_vec b__0 (Ox"3B9" : mword 12) then
+ read_reg pmp9cfg_ref >>= fun w__72 : Pmpcfg_ent =>
+ read_reg pmp10cfg_ref >>= fun w__73 : Pmpcfg_ent =>
+ (pmpTORLocked w__73) >>= fun w__74 : bool =>
+ ((read_reg pmpaddr9_ref) : M (mword 64)) >>= fun w__75 : mword 64 =>
+ write_reg pmpaddr9_ref (pmpWriteAddr (pmpLocked w__72) w__74 w__75 value) >>
+ ((read_reg pmpaddr9_ref) : M (mword 64)) >>= fun w__76 : mword 64 => returnm (Some w__76)
+ else if eq_vec b__0 (Ox"3BA" : mword 12) then
+ read_reg pmp10cfg_ref >>= fun w__77 : Pmpcfg_ent =>
+ read_reg pmp11cfg_ref >>= fun w__78 : Pmpcfg_ent =>
+ (pmpTORLocked w__78) >>= fun w__79 : bool =>
+ ((read_reg pmpaddr10_ref) : M (mword 64)) >>= fun w__80 : mword 64 =>
+ write_reg pmpaddr10_ref (pmpWriteAddr (pmpLocked w__77) w__79 w__80 value) >>
+ ((read_reg pmpaddr10_ref) : M (mword 64)) >>= fun w__81 : mword 64 =>
+ returnm (Some w__81)
+ else if eq_vec b__0 (Ox"3BB" : mword 12) then
+ read_reg pmp11cfg_ref >>= fun w__82 : Pmpcfg_ent =>
+ read_reg pmp12cfg_ref >>= fun w__83 : Pmpcfg_ent =>
+ (pmpTORLocked w__83) >>= fun w__84 : bool =>
+ ((read_reg pmpaddr11_ref) : M (mword 64)) >>= fun w__85 : mword 64 =>
+ write_reg pmpaddr11_ref (pmpWriteAddr (pmpLocked w__82) w__84 w__85 value) >>
+ ((read_reg pmpaddr11_ref) : M (mword 64)) >>= fun w__86 : mword 64 =>
+ returnm (Some w__86)
+ else if eq_vec b__0 (Ox"3BC" : mword 12) then
+ read_reg pmp12cfg_ref >>= fun w__87 : Pmpcfg_ent =>
+ read_reg pmp13cfg_ref >>= fun w__88 : Pmpcfg_ent =>
+ (pmpTORLocked w__88) >>= fun w__89 : bool =>
+ ((read_reg pmpaddr12_ref) : M (mword 64)) >>= fun w__90 : mword 64 =>
+ write_reg pmpaddr12_ref (pmpWriteAddr (pmpLocked w__87) w__89 w__90 value) >>
+ ((read_reg pmpaddr12_ref) : M (mword 64)) >>= fun w__91 : mword 64 =>
+ returnm (Some w__91)
+ else if eq_vec b__0 (Ox"3BD" : mword 12) then
+ read_reg pmp13cfg_ref >>= fun w__92 : Pmpcfg_ent =>
+ read_reg pmp14cfg_ref >>= fun w__93 : Pmpcfg_ent =>
+ (pmpTORLocked w__93) >>= fun w__94 : bool =>
+ ((read_reg pmpaddr13_ref) : M (mword 64)) >>= fun w__95 : mword 64 =>
+ write_reg pmpaddr13_ref (pmpWriteAddr (pmpLocked w__92) w__94 w__95 value) >>
+ ((read_reg pmpaddr13_ref) : M (mword 64)) >>= fun w__96 : mword 64 =>
+ returnm (Some w__96)
+ else if eq_vec b__0 (Ox"3BE" : mword 12) then
+ read_reg pmp14cfg_ref >>= fun w__97 : Pmpcfg_ent =>
+ read_reg pmp15cfg_ref >>= fun w__98 : Pmpcfg_ent =>
+ (pmpTORLocked w__98) >>= fun w__99 : bool =>
+ ((read_reg pmpaddr14_ref) : M (mword 64)) >>= fun w__100 : mword 64 =>
+ write_reg pmpaddr14_ref (pmpWriteAddr (pmpLocked w__97) w__99 w__100 value) >>
+ ((read_reg pmpaddr14_ref) : M (mword 64)) >>= fun w__101 : mword 64 =>
+ returnm (Some w__101)
+ else if eq_vec b__0 (Ox"3BF" : mword 12) then
+ read_reg pmp15cfg_ref >>= fun w__102 : Pmpcfg_ent =>
+ ((read_reg pmpaddr15_ref) : M (mword 64)) >>= fun w__103 : mword 64 =>
+ write_reg pmpaddr15_ref (pmpWriteAddr (pmpLocked w__102) false w__103 value) >>
+ ((read_reg pmpaddr15_ref) : M (mword 64)) >>= fun w__104 : mword 64 =>
+ returnm (Some w__104)
+ else if eq_vec b__0 (Ox"B00" : mword 12) then
+ ((read_reg mcycle_ref) : M (mword 64)) >>= fun w__105 : mword 64 =>
+ write_reg mcycle_ref (update_subrange_vec_dec w__105 (Z.sub 64 1) 0 value) >>
+ returnm (Some value)
+ else if eq_vec b__0 (Ox"B02" : mword 12) then
+ ((read_reg minstret_ref) : M (mword 64)) >>= fun w__106 : mword 64 =>
+ write_reg minstret_ref (update_subrange_vec_dec w__106 (Z.sub 64 1) 0 value) >>
+ write_reg minstret_written_ref true >> returnm (Some value)
+ else if eq_vec b__0 (Ox"7A0" : mword 12) then
write_reg tselect_ref value >>
- ((read_reg tselect_ref) : M (mword 64)) >>= fun w__74 : mword 64 =>
- returnm ((Some
- (w__74))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- read_reg mstatus_ref >>= fun w__75 : Mstatus =>
- (legalize_sstatus w__75 value) >>= fun w__76 : Mstatus =>
- write_reg mstatus_ref w__76 >>
- read_reg mstatus_ref >>= fun w__77 : Mstatus =>
- returnm ((Some
- (_get_Mstatus_bits w__77))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0] : mword 12))) then
- read_reg sedeleg_ref >>= fun w__78 : Sedeleg =>
- write_reg sedeleg_ref (legalize_sedeleg w__78 value) >>
- read_reg sedeleg_ref >>= fun w__79 : Sedeleg =>
- returnm ((Some
- (_get_Sedeleg_bits w__79))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ ((read_reg tselect_ref) : M (mword 64)) >>= fun w__107 : mword 64 =>
+ returnm (Some w__107)
+ else if eq_vec b__0 (Ox"100" : mword 12) then
+ read_reg mstatus_ref >>= fun w__108 : Mstatus =>
+ (legalize_sstatus w__108 value) >>= fun w__109 : Mstatus =>
+ write_reg mstatus_ref w__109 >>
+ read_reg mstatus_ref >>= fun w__110 : Mstatus => returnm (Some (_get_Mstatus_bits w__110))
+ else if eq_vec b__0 (Ox"102" : mword 12) then
+ read_reg sedeleg_ref >>= fun w__111 : Sedeleg =>
+ write_reg sedeleg_ref (legalize_sedeleg w__111 value) >>
+ read_reg sedeleg_ref >>= fun w__112 : Sedeleg => returnm (Some (_get_Sedeleg_bits w__112))
+ else if eq_vec b__0 (Ox"103" : mword 12) then
(_set_Sinterrupts_bits sideleg_ref value) >>
- read_reg sideleg_ref >>= fun w__80 : Sinterrupts =>
- returnm ((Some
- (_get_Sinterrupts_bits w__80))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B0] : mword 12))) then
- read_reg mie_ref >>= fun w__81 : Minterrupts =>
- read_reg mideleg_ref >>= fun w__82 : Minterrupts =>
- (legalize_sie w__81 w__82 value) >>= fun w__83 : Minterrupts =>
- write_reg mie_ref w__83 >>
- read_reg mie_ref >>= fun w__84 : Minterrupts =>
- returnm ((Some
- (_get_Minterrupts_bits w__84))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12))) then
- (set_stvec value) >>= fun w__85 : mword 64 =>
- returnm ((Some
- (w__85))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B1;B0] : mword 12))) then
- read_reg scounteren_ref >>= fun w__86 : Counteren =>
- (legalize_scounteren w__86 value) >>= fun w__87 : Counteren =>
- write_reg scounteren_ref w__87 >>
- read_reg scounteren_ref >>= fun w__88 : Counteren =>
- returnm ((Some
- (EXTZ 64 (_get_Counteren_bits w__88)))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0] : mword 12))) then
+ read_reg sideleg_ref >>= fun w__113 : Sinterrupts =>
+ returnm (Some (_get_Sinterrupts_bits w__113))
+ else if eq_vec b__0 (Ox"104" : mword 12) then
+ read_reg mie_ref >>= fun w__114 : Minterrupts =>
+ read_reg mideleg_ref >>= fun w__115 : Minterrupts =>
+ (legalize_sie w__114 w__115 value) >>= fun w__116 : Minterrupts =>
+ write_reg mie_ref w__116 >>
+ read_reg mie_ref >>= fun w__117 : Minterrupts =>
+ returnm (Some (_get_Minterrupts_bits w__117))
+ else if eq_vec b__0 (Ox"105" : mword 12) then
+ (set_stvec value) >>= fun w__118 : mword 64 => returnm (Some w__118)
+ else if eq_vec b__0 (Ox"106" : mword 12) then
+ read_reg scounteren_ref >>= fun w__119 : Counteren =>
+ write_reg scounteren_ref (legalize_scounteren w__119 value) >>
+ read_reg scounteren_ref >>= fun w__120 : Counteren =>
+ returnm (Some (EXTZ 64 (_get_Counteren_bits w__120)))
+ else if eq_vec b__0 (Ox"140" : mword 12) then
write_reg sscratch_ref value >>
- ((read_reg sscratch_ref) : M (mword 64)) >>= fun w__89 : mword 64 =>
- returnm ((Some
- (w__89))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B1] : mword 12))) then
- (set_xret_target Supervisor value) >>= fun w__90 : mword 64 =>
- returnm ((Some
- (w__90))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B0] : mword 12))) then
+ ((read_reg sscratch_ref) : M (mword 64)) >>= fun w__121 : mword 64 =>
+ returnm (Some w__121)
+ else if eq_vec b__0 (Ox"141" : mword 12) then
+ (set_xret_target Supervisor value) >>= fun w__122 : mword 64 => returnm (Some w__122)
+ else if eq_vec b__0 (Ox"142" : mword 12) then
(_set_Mcause_bits scause_ref value) >>
- read_reg scause_ref >>= fun w__91 : Mcause =>
- returnm ((Some
- (_get_Mcause_bits w__91))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B1;B1] : mword 12))) then
+ read_reg scause_ref >>= fun w__123 : Mcause => returnm (Some (_get_Mcause_bits w__123))
+ else if eq_vec b__0 (Ox"143" : mword 12) then
write_reg stval_ref value >>
- ((read_reg stval_ref) : M (mword 64)) >>= fun w__92 : mword 64 =>
- returnm ((Some
- (w__92))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B0;B1;B0;B0;B0;B1;B0;B0] : mword 12))) then
- read_reg mip_ref >>= fun w__93 : Minterrupts =>
- read_reg mideleg_ref >>= fun w__94 : Minterrupts =>
- (legalize_sip w__93 w__94 value) >>= fun w__95 : Minterrupts =>
- write_reg mip_ref w__95 >>
- read_reg mip_ref >>= fun w__96 : Minterrupts =>
- returnm ((Some
- (_get_Minterrupts_bits w__96))
- : option (mword 64))
- else if ((eq_vec b__0 (vec_of_bits [B0;B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B0] : mword 12))) then
- (cur_Architecture tt) >>= fun w__97 : Architecture =>
- ((read_reg satp_ref) : M (mword 64)) >>= fun w__98 : mword 64 =>
- write_reg satp_ref (legalize_satp w__97 w__98 value) >>
- ((read_reg satp_ref) : M (mword 64)) >>= fun w__99 : mword 64 =>
- returnm ((Some
- (w__99))
- : option (mword 64))
- else returnm (None : option (mword 64)))
+ ((read_reg stval_ref) : M (mword 64)) >>= fun w__124 : mword 64 => returnm (Some w__124)
+ else if eq_vec b__0 (Ox"144" : mword 12) then
+ read_reg mip_ref >>= fun w__125 : Minterrupts =>
+ read_reg mideleg_ref >>= fun w__126 : Minterrupts =>
+ (legalize_sip w__125 w__126 value) >>= fun w__127 : Minterrupts =>
+ write_reg mip_ref w__127 >>
+ read_reg mip_ref >>= fun w__128 : Minterrupts =>
+ returnm (Some (_get_Minterrupts_bits w__128))
+ else if eq_vec b__0 (Ox"180" : mword 12) then
+ (cur_Architecture tt) >>= fun w__129 : Architecture =>
+ ((read_reg satp_ref) : M (mword 64)) >>= fun w__130 : mword 64 =>
+ write_reg satp_ref (legalize_satp w__129 w__130 value) >>
+ ((read_reg satp_ref) : M (mword 64)) >>= fun w__131 : mword 64 => returnm (Some w__131)
+ else (ext_write_CSR csr value) : M (option (mword 64)))
: M (option (mword 64))
end) >>= fun res : option xlenbits =>
- (match res with
- | Some (v) =>
- returnm ((if ((get_config_print_reg tt)) then
- print_endline
- (String.append "CSR "
- (String.append (csr_name csr)
- (String.append " <- "
- (String.append (string_of_bits v)
- (String.append " (input: "
- (String.append (string_of_bits value) ")"))))))
- else tt)
- : unit)
- | None =>
- (ext_write_CSR csr value) >>= fun w__145 : bool =>
- returnm ((if sumbool_of_bool (w__145) then tt
- else print_bits "unhandled write to CSR " csr)
- : unit)
- end)
- : M (unit).
+ returnm (match res with
+ | Some v =>
+ if get_config_print_reg tt then
+ print_endline
+ (String.append "CSR "
+ (String.append (csr_name csr)
+ (String.append " <- "
+ (String.append (string_of_bits v)
+ (String.append " (input: " (String.append (string_of_bits value) ")"))))))
+ else tt
+ | None => print_bits "unhandled write to CSR " csr
+ end).
-Definition maybe_i_forwards (arg_ : bool)
-: string :=
-
+Definition maybe_i_forwards (arg_ : bool) : string :=
match arg_ with | true => "i" | false => "" end.
-Definition maybe_i_backwards (arg_ : string)
-: M (bool) :=
-
+Definition maybe_i_backwards (arg_ : string) : M (bool) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "i")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq p0_ "")) then
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))}))
+ (if generic_eq p0_ "i" then returnm true
+ else if generic_eq p0_ "" then returnm false
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (bool).
-Definition maybe_i_forwards_matches (arg_ : bool)
-: bool :=
-
+Definition maybe_i_forwards_matches (arg_ : bool) : bool :=
match arg_ with | true => true | false => true end.
-Definition maybe_i_backwards_matches (arg_ : string)
-: bool :=
-
+Definition maybe_i_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "i")) then true
- else if ((generic_eq p0_ "")) then true
+ if generic_eq p0_ "i" then true
+ else if generic_eq p0_ "" then true
else false.
-Definition _s797_ (_s798_ : string)
-: option string :=
-
- let _s799_ := _s798_ in
- if ((string_startswith _s799_ "")) then
- match (string_drop _s799_ (projT1 (string_length ""))) with | s_ => Some (s_) end
+Definition _s1058_ (_s1059_ : string) : option string :=
+ let _s1060_ := _s1059_ in
+ if string_startswith _s1060_ "" then
+ match (string_drop _s1060_ (projT1 (string_length ""))) with | s_ => Some s_ end
else None.
-Definition _s793_ (_s794_ : string)
-: option string :=
-
- let _s795_ := _s794_ in
- if ((string_startswith _s795_ "i")) then
- match (string_drop _s795_ (projT1 (string_length "i"))) with | s_ => Some (s_) end
- else None.
-
-Definition maybe_i_matches_prefix (arg_ : string)
-: M (option ((bool * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s796_ := arg_ in
- (if ((match (_s793_ _s796_) with | Some (s_) => true | _ => false end)) then
- (match (_s793_ _s796_) with
- | Some (s_) =>
- returnm ((Some
- ((true, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s797_ _s796_) with | Some (s_) => true | _ => false end)) then
- (match (_s797_ _s796_) with
- | Some (s_) =>
- returnm ((Some
- ((false, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((bool * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
+Definition _s1054_ (_s1055_ : string) : option string :=
+ let _s1056_ := _s1055_ in
+ if string_startswith _s1056_ "i" then
+ match (string_drop _s1056_ (projT1 (string_length "i"))) with | s_ => Some s_ end
+ else None.
+
+Definition maybe_i_matches_prefix (arg_ : string)
+: M (option ((bool * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1057_ := arg_ in
+ (if match (_s1054_ _s1057_) with | Some s_ => true | _ => false end then
+ (match (_s1054_ _s1057_) with
+ | Some s_ =>
+ returnm (Some
+ (true, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((bool * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((bool * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((bool * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1058_ _s1057_) with | Some s_ => true | _ => false end then
+ (match (_s1058_ _s1057_) with
+ | Some s_ =>
+ returnm (Some
+ (false, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((bool * {n : Z & ArithFact (n >=? 0)}))).
-Definition csr_mnemonic_forwards (arg_ : csrop)
-: string :=
-
+Definition csr_mnemonic_forwards (arg_ : csrop) : string :=
match arg_ with | CSRRW => "csrrw" | CSRRS => "csrrs" | CSRRC => "csrrc" end.
-Definition csr_mnemonic_backwards (arg_ : string)
-: M (csrop) :=
-
+Definition csr_mnemonic_backwards (arg_ : string) : M (csrop) :=
let p0_ := arg_ in
- (if ((generic_eq p0_ "csrrw")) then returnm (CSRRW : csrop)
- else if ((generic_eq p0_ "csrrs")) then returnm (CSRRS : csrop)
- else if ((generic_eq p0_ "csrrc")) then returnm (CSRRC : csrop)
+ (if generic_eq p0_ "csrrw" then returnm CSRRW
+ else if generic_eq p0_ "csrrs" then returnm CSRRS
+ else if generic_eq p0_ "csrrc" then returnm CSRRC
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (csrop).
-Definition csr_mnemonic_forwards_matches (arg_ : csrop)
-: bool :=
-
+Definition csr_mnemonic_forwards_matches (arg_ : csrop) : bool :=
match arg_ with | CSRRW => true | CSRRS => true | CSRRC => true end.
-Definition csr_mnemonic_backwards_matches (arg_ : string)
-: bool :=
-
+Definition csr_mnemonic_backwards_matches (arg_ : string) : bool :=
let p0_ := arg_ in
- if ((generic_eq p0_ "csrrw")) then true
- else if ((generic_eq p0_ "csrrs")) then true
- else if ((generic_eq p0_ "csrrc")) then true
+ if generic_eq p0_ "csrrw" then true
+ else if generic_eq p0_ "csrrs" then true
+ else if generic_eq p0_ "csrrc" then true
else false.
-Definition _s809_ (_s810_ : string)
-: option string :=
-
- let _s811_ := _s810_ in
- if ((string_startswith _s811_ "csrrc")) then
- match (string_drop _s811_ (projT1 (string_length "csrrc"))) with | s_ => Some (s_) end
+Definition _s1070_ (_s1071_ : string) : option string :=
+ let _s1072_ := _s1071_ in
+ if string_startswith _s1072_ "csrrc" then
+ match (string_drop _s1072_ (projT1 (string_length "csrrc"))) with | s_ => Some s_ end
else None.
-Definition _s805_ (_s806_ : string)
-: option string :=
-
- let _s807_ := _s806_ in
- if ((string_startswith _s807_ "csrrs")) then
- match (string_drop _s807_ (projT1 (string_length "csrrs"))) with | s_ => Some (s_) end
+Definition _s1066_ (_s1067_ : string) : option string :=
+ let _s1068_ := _s1067_ in
+ if string_startswith _s1068_ "csrrs" then
+ match (string_drop _s1068_ (projT1 (string_length "csrrs"))) with | s_ => Some s_ end
else None.
-Definition _s801_ (_s802_ : string)
-: option string :=
-
- let _s803_ := _s802_ in
- if ((string_startswith _s803_ "csrrw")) then
- match (string_drop _s803_ (projT1 (string_length "csrrw"))) with | s_ => Some (s_) end
- else None.
-
-Definition csr_mnemonic_matches_prefix (arg_ : string)
-: M (option ((csrop * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s804_ := arg_ in
- (if ((match (_s801_ _s804_) with | Some (s_) => true | _ => false end)) then
- (match (_s801_ _s804_) with
- | Some (s_) =>
- returnm ((Some
- ((CSRRW, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((csrop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((csrop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((csrop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s805_ _s804_) with | Some (s_) => true | _ => false end)) then
- (match (_s805_ _s804_) with
- | Some (s_) =>
- returnm ((Some
- ((CSRRS, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((csrop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((csrop * {n : Z & ArithFact (n >= 0)})))
- end)
- : M (option ((csrop * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s809_ _s804_) with | Some (s_) => true | _ => false end)) then
- (match (_s809_ _s804_) with
- | Some (s_) =>
- returnm ((Some
- ((CSRRC, build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((csrop * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((csrop * {n : Z & ArithFact (n >= 0)})))
+Definition _s1062_ (_s1063_ : string) : option string :=
+ let _s1064_ := _s1063_ in
+ if string_startswith _s1064_ "csrrw" then
+ match (string_drop _s1064_ (projT1 (string_length "csrrw"))) with | s_ => Some s_ end
+ else None.
+
+Definition csr_mnemonic_matches_prefix (arg_ : string)
+: M (option ((csrop * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1065_ := arg_ in
+ (if match (_s1062_ _s1065_) with | Some s_ => true | _ => false end then
+ (match (_s1062_ _s1065_) with
+ | Some s_ =>
+ returnm (Some
+ (CSRRW, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((csrop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((csrop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1066_ _s1065_) with | Some s_ => true | _ => false end then
+ (match (_s1066_ _s1065_) with
+ | Some s_ =>
+ returnm (Some
+ (CSRRS, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((csrop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((csrop * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1070_ _s1065_) with | Some s_ => true | _ => false end then
+ (match (_s1070_ _s1065_) with
+ | Some s_ =>
+ returnm (Some
+ (CSRRC, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((csrop * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((csrop * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((csrop * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition encdec_rounding_mode_forwards (arg_ : rounding_mode) : mword 3 :=
+ match arg_ with
+ | RM_RNE => 'b"000" : mword 3
+ | RM_RTZ => 'b"001" : mword 3
+ | RM_RDN => 'b"010" : mword 3
+ | RM_RUP => 'b"011" : mword 3
+ | RM_RMM => 'b"100" : mword 3
+ | RM_DYN => 'b"111" : mword 3
+ end.
+
+Definition encdec_rounding_mode_backwards (arg_ : mword 3) : M (rounding_mode) :=
+ let b__0 := arg_ in
+ (if eq_vec b__0 ('b"000" : mword 3) then returnm RM_RNE
+ else if eq_vec b__0 ('b"001" : mword 3) then returnm RM_RTZ
+ else if eq_vec b__0 ('b"010" : mword 3) then returnm RM_RDN
+ else if eq_vec b__0 ('b"011" : mword 3) then returnm RM_RUP
+ else if eq_vec b__0 ('b"100" : mword 3) then returnm RM_RMM
+ else if eq_vec b__0 ('b"111" : mword 3) then returnm RM_DYN
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (rounding_mode).
+
+Definition encdec_rounding_mode_forwards_matches (arg_ : rounding_mode) : bool :=
+ match arg_ with
+ | RM_RNE => true
+ | RM_RTZ => true
+ | RM_RDN => true
+ | RM_RUP => true
+ | RM_RMM => true
+ | RM_DYN => true
+ end.
+
+Definition encdec_rounding_mode_backwards_matches (arg_ : mword 3) : bool :=
+ let b__0 := arg_ in
+ if eq_vec b__0 ('b"000" : mword 3) then true
+ else if eq_vec b__0 ('b"001" : mword 3) then true
+ else if eq_vec b__0 ('b"010" : mword 3) then true
+ else if eq_vec b__0 ('b"011" : mword 3) then true
+ else if eq_vec b__0 ('b"100" : mword 3) then true
+ else if eq_vec b__0 ('b"111" : mword 3) then true
+ else false.
+
+Definition frm_mnemonic_forwards (arg_ : rounding_mode) : string :=
+ match arg_ with
+ | RM_RNE => "rne"
+ | RM_RTZ => "rtz"
+ | RM_RDN => "rdn"
+ | RM_RUP => "rup"
+ | RM_RMM => "rmm"
+ | RM_DYN => "dyn"
+ end.
+
+Definition frm_mnemonic_backwards (arg_ : string) : M (rounding_mode) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "rne" then returnm RM_RNE
+ else if generic_eq p0_ "rtz" then returnm RM_RTZ
+ else if generic_eq p0_ "rdn" then returnm RM_RDN
+ else if generic_eq p0_ "rup" then returnm RM_RUP
+ else if generic_eq p0_ "rmm" then returnm RM_RMM
+ else if generic_eq p0_ "dyn" then returnm RM_DYN
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (rounding_mode).
+
+Definition frm_mnemonic_forwards_matches (arg_ : rounding_mode) : bool :=
+ match arg_ with
+ | RM_RNE => true
+ | RM_RTZ => true
+ | RM_RDN => true
+ | RM_RUP => true
+ | RM_RMM => true
+ | RM_DYN => true
+ end.
+
+Definition frm_mnemonic_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "rne" then true
+ else if generic_eq p0_ "rtz" then true
+ else if generic_eq p0_ "rdn" then true
+ else if generic_eq p0_ "rup" then true
+ else if generic_eq p0_ "rmm" then true
+ else if generic_eq p0_ "dyn" then true
+ else false.
+
+Definition _s1094_ (_s1095_ : string) : option string :=
+ let _s1096_ := _s1095_ in
+ if string_startswith _s1096_ "dyn" then
+ match (string_drop _s1096_ (projT1 (string_length "dyn"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1090_ (_s1091_ : string) : option string :=
+ let _s1092_ := _s1091_ in
+ if string_startswith _s1092_ "rmm" then
+ match (string_drop _s1092_ (projT1 (string_length "rmm"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1086_ (_s1087_ : string) : option string :=
+ let _s1088_ := _s1087_ in
+ if string_startswith _s1088_ "rup" then
+ match (string_drop _s1088_ (projT1 (string_length "rup"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1082_ (_s1083_ : string) : option string :=
+ let _s1084_ := _s1083_ in
+ if string_startswith _s1084_ "rdn" then
+ match (string_drop _s1084_ (projT1 (string_length "rdn"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1078_ (_s1079_ : string) : option string :=
+ let _s1080_ := _s1079_ in
+ if string_startswith _s1080_ "rtz" then
+ match (string_drop _s1080_ (projT1 (string_length "rtz"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1074_ (_s1075_ : string) : option string :=
+ let _s1076_ := _s1075_ in
+ if string_startswith _s1076_ "rne" then
+ match (string_drop _s1076_ (projT1 (string_length "rne"))) with | s_ => Some s_ end
+ else None.
+
+Definition frm_mnemonic_matches_prefix (arg_ : string)
+: M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1077_ := arg_ in
+ (if match (_s1074_ _s1077_) with | Some s_ => true | _ => false end then
+ (match (_s1074_ _s1077_) with
+ | Some s_ =>
+ returnm (Some
+ (RM_RNE, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1078_ _s1077_) with | Some s_ => true | _ => false end then
+ (match (_s1078_ _s1077_) with
+ | Some s_ =>
+ returnm (Some
+ (RM_RTZ, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1082_ _s1077_) with | Some s_ => true | _ => false end then
+ (match (_s1082_ _s1077_) with
+ | Some s_ =>
+ returnm (Some
+ (RM_RDN, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1086_ _s1077_) with | Some s_ => true | _ => false end then
+ (match (_s1086_ _s1077_) with
+ | Some s_ =>
+ returnm (Some
+ (RM_RUP, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1090_ _s1077_) with | Some s_ => true | _ => false end then
+ (match (_s1090_ _s1077_) with
+ | Some s_ =>
+ returnm (Some
+ (RM_RMM, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1094_ _s1077_) with | Some s_ => true | _ => false end then
+ (match (_s1094_ _s1077_) with
+ | Some s_ =>
+ returnm (Some
+ (RM_DYN, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((rounding_mode * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition select_instr_or_fcsr_rm (instr_rm : rounding_mode) : M (rounding_mode) :=
+ (if generic_eq instr_rm RM_DYN then
+ read_reg fcsr_ref >>= fun w__0 : Fcsr =>
+ (encdec_rounding_mode_backwards (_get_Fcsr_FRM w__0))
+ : M (rounding_mode)
+ else returnm instr_rm)
+ : M (rounding_mode).
+
+Definition nxFlag '(tt : unit) : mword 5 := 'b"00001" : mword 5.
+
+Definition ufFlag '(tt : unit) : mword 5 := 'b"00010" : mword 5.
+
+Definition ofFlag '(tt : unit) : mword 5 := 'b"00100" : mword 5.
+
+Definition dzFlag '(tt : unit) : mword 5 := 'b"01000" : mword 5.
+
+Definition nvFlag '(tt : unit) : mword 5 := 'b"10000" : mword 5.
+
+Definition fsplit_S (x32 : mword 32) : (mword 1 * mword 8 * mword 23) :=
+ (subrange_vec_dec x32 31 31, subrange_vec_dec x32 30 23, subrange_vec_dec x32 22 0).
+
+Definition fmake_S (sign : mword 1) (exp : mword 8) (mant : mword 23) : mword 32 :=
+ concat_vec sign (concat_vec exp mant).
+
+Definition canonical_NaN_S '(tt : unit) : mword 32 := Ox"7FC00000" : mword 32.
+
+Definition f_is_neg_inf_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec sign ('b"1" : mword 1))
+ (andb (eq_vec exp (ones 8)) (eq_vec mant (zeros_implicit 23))).
+
+Definition f_is_neg_norm_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec sign ('b"1" : mword 1))
+ (andb (neq_vec exp (zeros_implicit 8)) (neq_vec exp (ones 8))).
+
+Definition f_is_neg_subnorm_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec sign ('b"1" : mword 1))
+ (andb (eq_vec exp (zeros_implicit 8)) (neq_vec mant (zeros_implicit 23))).
+
+Definition f_is_neg_zero_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec sign (ones 1))
+ (andb (eq_vec exp (zeros_implicit 8)) (eq_vec mant (zeros_implicit 23))).
+
+Definition f_is_pos_zero_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec sign (zeros_implicit 1))
+ (andb (eq_vec exp (zeros_implicit 8)) (eq_vec mant (zeros_implicit 23))).
+
+Definition f_is_pos_subnorm_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec sign (zeros_implicit 1))
+ (andb (eq_vec exp (zeros_implicit 8)) (neq_vec mant (zeros_implicit 23))).
+
+Definition f_is_pos_norm_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec sign (zeros_implicit 1))
+ (andb (neq_vec exp (zeros_implicit 8)) (neq_vec exp (ones 8))).
+
+Definition f_is_pos_inf_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec sign (zeros_implicit 1))
+ (andb (eq_vec exp (ones 8)) (eq_vec mant (zeros_implicit 23))).
+
+Definition f_is_SNaN_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec exp (ones 8))
+ (andb (eq_bit (access_vec_dec mant 22) B0) (neq_vec mant (zeros_implicit 23))).
+
+Definition f_is_QNaN_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec exp (ones 8)) (eq_bit (access_vec_dec mant 22) B1).
+
+Definition f_is_NaN_S (x32 : mword 32) : bool :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ andb (eq_vec exp (ones 8)) (neq_vec mant (zeros_implicit 23)).
+
+Definition negate_S (x32 : mword 32) : mword 32 :=
+ let '(sign, exp, mant) := fsplit_S x32 in
+ let new_sign := if eq_vec sign ('b"0" : mword 1) then 'b"1" : mword 1 else 'b"0" : mword 1 in
+ fmake_S new_sign exp mant.
+
+Definition feq_quiet_S (v1 : mword 32) (v2 : mword 32) : (bool * mword 5) :=
+ let '(s1, e1, m1) := fsplit_S v1 in
+ let '(s2, e2, m2) := fsplit_S v2 in
+ let v1Is0 := orb (f_is_neg_zero_S v1) (f_is_pos_zero_S v1) in
+ let v2Is0 := orb (f_is_neg_zero_S v2) (f_is_pos_zero_S v2) in
+ let result := orb (eq_vec v1 v2) (andb v1Is0 v2Is0) in
+ let fflags := if orb (f_is_SNaN_S v1) (f_is_SNaN_S v2) then nvFlag tt else zeros_implicit 5 in
+ (result, fflags).
+
+Definition flt_S (v1 : mword 32) (v2 : mword 32) (is_quiet : bool) : (bool * mword 5) :=
+ let '(s1, e1, m1) := fsplit_S v1 in
+ let '(s2, e2, m2) := fsplit_S v2 in
+ let result : bool :=
+ if andb (eq_vec s1 ('b"0" : mword 1)) (eq_vec s2 ('b"0" : mword 1)) then
+ if eq_vec e1 e2 then Z.ltb (projT1 (uint m1)) (projT1 (uint m2))
+ else Z.ltb (projT1 (uint e1)) (projT1 (uint e2))
+ else if andb (eq_vec s1 ('b"0" : mword 1)) (eq_vec s2 ('b"1" : mword 1)) then false
+ else if andb (eq_vec s1 ('b"1" : mword 1)) (eq_vec s2 ('b"0" : mword 1)) then true
+ else if eq_vec e1 e2 then Z.gtb (projT1 (uint m1)) (projT1 (uint m2))
+ else Z.gtb (projT1 (uint e1)) (projT1 (uint e2)) in
+ let fflags :=
+ if sumbool_of_bool is_quiet then
+ if orb (f_is_SNaN_S v1) (f_is_SNaN_S v2) then nvFlag tt
+ else zeros_implicit 5
+ else if orb (f_is_NaN_S v1) (f_is_NaN_S v2) then nvFlag tt
+ else zeros_implicit 5 in
+ (result, fflags).
+
+Definition fle_S (v1 : mword 32) (v2 : mword 32) (is_quiet : bool) : (bool * mword 5) :=
+ let '(s1, e1, m1) := fsplit_S v1 in
+ let '(s2, e2, m2) := fsplit_S v2 in
+ let v1Is0 := orb (f_is_neg_zero_S v1) (f_is_pos_zero_S v1) in
+ let v2Is0 := orb (f_is_neg_zero_S v2) (f_is_pos_zero_S v2) in
+ let result : bool :=
+ if andb (eq_vec s1 ('b"0" : mword 1)) (eq_vec s2 ('b"0" : mword 1)) then
+ if eq_vec e1 e2 then Z.leb (projT1 (uint m1)) (projT1 (uint m2))
+ else Z.ltb (projT1 (uint e1)) (projT1 (uint e2))
+ else if andb (eq_vec s1 ('b"0" : mword 1)) (eq_vec s2 ('b"1" : mword 1)) then
+ andb v1Is0 v2Is0
+ else if andb (eq_vec s1 ('b"1" : mword 1)) (eq_vec s2 ('b"0" : mword 1)) then true
+ else if eq_vec e1 e2 then Z.geb (projT1 (uint m1)) (projT1 (uint m2))
+ else Z.gtb (projT1 (uint e1)) (projT1 (uint e2)) in
+ let fflags :=
+ if sumbool_of_bool is_quiet then
+ if orb (f_is_SNaN_S v1) (f_is_SNaN_S v2) then nvFlag tt
+ else zeros_implicit 5
+ else if orb (f_is_NaN_S v1) (f_is_NaN_S v2) then nvFlag tt
+ else zeros_implicit 5 in
+ (result, fflags).
+
+Definition nan_box (val_32b : mword 32) : mword 64 := concat_vec (Ox"FFFFFFFF" : mword 32) val_32b.
+
+Definition nan_unbox (regval : mword 64) : mword 32 :=
+ if eq_vec (subrange_vec_dec regval 63 32) (Ox"FFFFFFFF" : mword (63 - 32 + 1)) then
+ subrange_vec_dec regval 31 0
+ else canonical_NaN_S tt.
+
+Definition is_RV32F_or_RV64F '(tt : unit) : M (bool) :=
+ projT1_m
+ ((and_boolMP (build_trivial_ex ((haveFExt tt) : M (bool)))
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && ((64 =? 32) || (64 =? 64))) _bool = true)})).
+
+Definition is_RV64F '(tt : unit) : M (bool) :=
+ projT1_m
+ ((and_boolMP (build_trivial_ex ((haveFExt tt) : M (bool)))
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 64)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 64) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 64) _bool)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 64)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 64)) _bool = true)})).
+
+Definition is_RV32D_or_RV64D '(tt : unit) : M (bool) :=
+ projT1_m
+ ((and_boolMP (build_trivial_ex ((haveDExt tt) : M (bool)))
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && ((64 =? 32) || (64 =? 64))) _bool = true)})).
+
+Definition is_RV64D '(tt : unit) : M (bool) :=
+ projT1_m
+ ((and_boolMP (build_trivial_ex ((haveDExt tt) : M (bool)))
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 64)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 64) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 64) _bool)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 64)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 64)) _bool = true)})).
+
+Definition process_fload64 (rd : mword 5) (addr : mword 64) (value : MemoryOpResult (mword 64))
+: M (Retired) :=
+ (match value with
+ | MemValue result => (wF_bits rd result) >> returnm RETIRE_SUCCESS
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
+ end)
+ : M (Retired).
+
+Definition process_fload32 (rd : mword 5) (addr : mword 64) (value : MemoryOpResult (mword 32))
+: M (Retired) :=
+ (match value with
+ | MemValue result => (wF_bits rd (nan_box result)) >> returnm RETIRE_SUCCESS
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
+ end)
+ : M (Retired).
+
+Definition process_fstore (vaddr : mword 64) (value : MemoryOpResult bool) : M (Retired) :=
+ (match value with
+ | MemValue true => returnm RETIRE_SUCCESS
+ | MemValue false => (internal_error "store got false from mem_write_value") : M (Retired)
+ | MemException e => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ end)
+ : M (Retired).
+
+Definition f_madd_type_mnemonic_S_forwards (arg_ : f_madd_op_S) : string :=
+ match arg_ with
+ | FMADD_S => "fmadd.s"
+ | FMSUB_S => "fmsub.s"
+ | FNMSUB_S => "fnmsub.s"
+ | FNMADD_S => "fnmadd.s"
+ end.
+
+Definition f_madd_type_mnemonic_S_backwards (arg_ : string) : M (f_madd_op_S) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fmadd.s" then returnm FMADD_S
+ else if generic_eq p0_ "fmsub.s" then returnm FMSUB_S
+ else if generic_eq p0_ "fnmsub.s" then returnm FNMSUB_S
+ else if generic_eq p0_ "fnmadd.s" then returnm FNMADD_S
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_madd_op_S).
+
+Definition f_madd_type_mnemonic_S_forwards_matches (arg_ : f_madd_op_S) : bool :=
+ match arg_ with | FMADD_S => true | FMSUB_S => true | FNMSUB_S => true | FNMADD_S => true end.
+
+Definition f_madd_type_mnemonic_S_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fmadd.s" then true
+ else if generic_eq p0_ "fmsub.s" then true
+ else if generic_eq p0_ "fnmsub.s" then true
+ else if generic_eq p0_ "fnmadd.s" then true
+ else false.
+
+Definition _s1110_ (_s1111_ : string) : option string :=
+ let _s1112_ := _s1111_ in
+ if string_startswith _s1112_ "fnmadd.s" then
+ match (string_drop _s1112_ (projT1 (string_length "fnmadd.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1106_ (_s1107_ : string) : option string :=
+ let _s1108_ := _s1107_ in
+ if string_startswith _s1108_ "fnmsub.s" then
+ match (string_drop _s1108_ (projT1 (string_length "fnmsub.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1102_ (_s1103_ : string) : option string :=
+ let _s1104_ := _s1103_ in
+ if string_startswith _s1104_ "fmsub.s" then
+ match (string_drop _s1104_ (projT1 (string_length "fmsub.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1098_ (_s1099_ : string) : option string :=
+ let _s1100_ := _s1099_ in
+ if string_startswith _s1100_ "fmadd.s" then
+ match (string_drop _s1100_ (projT1 (string_length "fmadd.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_madd_type_mnemonic_S_matches_prefix (arg_ : string)
+: M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1101_ := arg_ in
+ (if match (_s1098_ _s1101_) with | Some s_ => true | _ => false end then
+ (match (_s1098_ _s1101_) with
+ | Some s_ =>
+ returnm (Some
+ (FMADD_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1102_ _s1101_) with | Some s_ => true | _ => false end then
+ (match (_s1102_ _s1101_) with
+ | Some s_ =>
+ returnm (Some
+ (FMSUB_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1106_ _s1101_) with | Some s_ => true | _ => false end then
+ (match (_s1106_ _s1101_) with
+ | Some s_ =>
+ returnm (Some
+ (FNMSUB_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1110_ _s1101_) with | Some s_ => true | _ => false end then
+ (match (_s1110_ _s1101_) with
+ | Some s_ =>
+ returnm (Some
+ (FNMADD_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_madd_op_S * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition f_bin_rm_type_mnemonic_S_forwards (arg_ : f_bin_rm_op_S) : string :=
+ match arg_ with
+ | FADD_S => "fadd.s"
+ | FSUB_S => "fsub.s"
+ | FMUL_S => "fmul.s"
+ | FDIV_S => "fdiv.s"
+ end.
+
+Definition f_bin_rm_type_mnemonic_S_backwards (arg_ : string) : M (f_bin_rm_op_S) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fadd.s" then returnm FADD_S
+ else if generic_eq p0_ "fsub.s" then returnm FSUB_S
+ else if generic_eq p0_ "fmul.s" then returnm FMUL_S
+ else if generic_eq p0_ "fdiv.s" then returnm FDIV_S
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_bin_rm_op_S).
+
+Definition f_bin_rm_type_mnemonic_S_forwards_matches (arg_ : f_bin_rm_op_S) : bool :=
+ match arg_ with | FADD_S => true | FSUB_S => true | FMUL_S => true | FDIV_S => true end.
+
+Definition f_bin_rm_type_mnemonic_S_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fadd.s" then true
+ else if generic_eq p0_ "fsub.s" then true
+ else if generic_eq p0_ "fmul.s" then true
+ else if generic_eq p0_ "fdiv.s" then true
+ else false.
+
+Definition _s1126_ (_s1127_ : string) : option string :=
+ let _s1128_ := _s1127_ in
+ if string_startswith _s1128_ "fdiv.s" then
+ match (string_drop _s1128_ (projT1 (string_length "fdiv.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1122_ (_s1123_ : string) : option string :=
+ let _s1124_ := _s1123_ in
+ if string_startswith _s1124_ "fmul.s" then
+ match (string_drop _s1124_ (projT1 (string_length "fmul.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1118_ (_s1119_ : string) : option string :=
+ let _s1120_ := _s1119_ in
+ if string_startswith _s1120_ "fsub.s" then
+ match (string_drop _s1120_ (projT1 (string_length "fsub.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1114_ (_s1115_ : string) : option string :=
+ let _s1116_ := _s1115_ in
+ if string_startswith _s1116_ "fadd.s" then
+ match (string_drop _s1116_ (projT1 (string_length "fadd.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_bin_rm_type_mnemonic_S_matches_prefix (arg_ : string)
+: M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1117_ := arg_ in
+ (if match (_s1114_ _s1117_) with | Some s_ => true | _ => false end then
+ (match (_s1114_ _s1117_) with
+ | Some s_ =>
+ returnm (Some
+ (FADD_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1118_ _s1117_) with | Some s_ => true | _ => false end then
+ (match (_s1118_ _s1117_) with
+ | Some s_ =>
+ returnm (Some
+ (FSUB_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1122_ _s1117_) with | Some s_ => true | _ => false end then
+ (match (_s1122_ _s1117_) with
+ | Some s_ =>
+ returnm (Some
+ (FMUL_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1126_ _s1117_) with | Some s_ => true | _ => false end then
+ (match (_s1126_ _s1117_) with
+ | Some s_ =>
+ returnm (Some
+ (FDIV_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition f_un_rm_type_mnemonic_S_forwards (arg_ : f_un_rm_op_S) : string :=
+ match arg_ with
+ | FSQRT_S => "fsqrt.s"
+ | FCVT_W_S => "fcvt.w.s"
+ | FCVT_WU_S => "fcvt.wu.s"
+ | FCVT_S_W => "fcvt.s.w"
+ | FCVT_S_WU => "fcvt.s.wu"
+ | FCVT_L_S => "fcvt.l.s"
+ | FCVT_LU_S => "fcvt.lu.s"
+ | FCVT_S_L => "fcvt.s.l"
+ | FCVT_S_LU => "fcvt.s.lu"
+ end.
+
+Definition f_un_rm_type_mnemonic_S_backwards (arg_ : string) : M (f_un_rm_op_S) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fsqrt.s" then returnm FSQRT_S
+ else if generic_eq p0_ "fcvt.w.s" then returnm FCVT_W_S
+ else if generic_eq p0_ "fcvt.wu.s" then returnm FCVT_WU_S
+ else if generic_eq p0_ "fcvt.s.w" then returnm FCVT_S_W
+ else if generic_eq p0_ "fcvt.s.wu" then returnm FCVT_S_WU
+ else if generic_eq p0_ "fcvt.l.s" then returnm FCVT_L_S
+ else if generic_eq p0_ "fcvt.lu.s" then returnm FCVT_LU_S
+ else if generic_eq p0_ "fcvt.s.l" then returnm FCVT_S_L
+ else if generic_eq p0_ "fcvt.s.lu" then returnm FCVT_S_LU
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_un_rm_op_S).
+
+Definition f_un_rm_type_mnemonic_S_forwards_matches (arg_ : f_un_rm_op_S) : bool :=
+ match arg_ with
+ | FSQRT_S => true
+ | FCVT_W_S => true
+ | FCVT_WU_S => true
+ | FCVT_S_W => true
+ | FCVT_S_WU => true
+ | FCVT_L_S => true
+ | FCVT_LU_S => true
+ | FCVT_S_L => true
+ | FCVT_S_LU => true
+ end.
+
+Definition f_un_rm_type_mnemonic_S_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fsqrt.s" then true
+ else if generic_eq p0_ "fcvt.w.s" then true
+ else if generic_eq p0_ "fcvt.wu.s" then true
+ else if generic_eq p0_ "fcvt.s.w" then true
+ else if generic_eq p0_ "fcvt.s.wu" then true
+ else if generic_eq p0_ "fcvt.l.s" then true
+ else if generic_eq p0_ "fcvt.lu.s" then true
+ else if generic_eq p0_ "fcvt.s.l" then true
+ else if generic_eq p0_ "fcvt.s.lu" then true
+ else false.
+
+Definition _s1162_ (_s1163_ : string) : option string :=
+ let _s1164_ := _s1163_ in
+ if string_startswith _s1164_ "fcvt.s.lu" then
+ match (string_drop _s1164_ (projT1 (string_length "fcvt.s.lu"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1158_ (_s1159_ : string) : option string :=
+ let _s1160_ := _s1159_ in
+ if string_startswith _s1160_ "fcvt.s.l" then
+ match (string_drop _s1160_ (projT1 (string_length "fcvt.s.l"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1154_ (_s1155_ : string) : option string :=
+ let _s1156_ := _s1155_ in
+ if string_startswith _s1156_ "fcvt.lu.s" then
+ match (string_drop _s1156_ (projT1 (string_length "fcvt.lu.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1150_ (_s1151_ : string) : option string :=
+ let _s1152_ := _s1151_ in
+ if string_startswith _s1152_ "fcvt.l.s" then
+ match (string_drop _s1152_ (projT1 (string_length "fcvt.l.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1146_ (_s1147_ : string) : option string :=
+ let _s1148_ := _s1147_ in
+ if string_startswith _s1148_ "fcvt.s.wu" then
+ match (string_drop _s1148_ (projT1 (string_length "fcvt.s.wu"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1142_ (_s1143_ : string) : option string :=
+ let _s1144_ := _s1143_ in
+ if string_startswith _s1144_ "fcvt.s.w" then
+ match (string_drop _s1144_ (projT1 (string_length "fcvt.s.w"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1138_ (_s1139_ : string) : option string :=
+ let _s1140_ := _s1139_ in
+ if string_startswith _s1140_ "fcvt.wu.s" then
+ match (string_drop _s1140_ (projT1 (string_length "fcvt.wu.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1134_ (_s1135_ : string) : option string :=
+ let _s1136_ := _s1135_ in
+ if string_startswith _s1136_ "fcvt.w.s" then
+ match (string_drop _s1136_ (projT1 (string_length "fcvt.w.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1130_ (_s1131_ : string) : option string :=
+ let _s1132_ := _s1131_ in
+ if string_startswith _s1132_ "fsqrt.s" then
+ match (string_drop _s1132_ (projT1 (string_length "fsqrt.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_un_rm_type_mnemonic_S_matches_prefix (arg_ : string)
+: M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1133_ := arg_ in
+ (if match (_s1130_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1130_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FSQRT_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1134_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1134_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_W_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1138_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1138_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_WU_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1142_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1142_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_S_W, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1146_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1146_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_S_WU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1150_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1150_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_L_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1154_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1154_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_LU_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1158_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1158_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_S_L, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1162_ _s1133_) with | Some s_ => true | _ => false end then
+ (match (_s1162_ _s1133_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_S_LU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_un_rm_op_S * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition f_bin_type_mnemonic_S_forwards (arg_ : f_bin_op_S) : string :=
+ match arg_ with
+ | FSGNJ_S => "fsgnj.s"
+ | FSGNJN_S => "fsgnjn.s"
+ | FSGNJX_S => "fsgnjx.s"
+ | FMIN_S => "fmin.s"
+ | FMAX_S => "fmax.s"
+ | FEQ_S => "feq.s"
+ | FLT_S => "flt.s"
+ | FLE_S => "fle.s"
+ end.
+
+Definition f_bin_type_mnemonic_S_backwards (arg_ : string) : M (f_bin_op_S) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fsgnj.s" then returnm FSGNJ_S
+ else if generic_eq p0_ "fsgnjn.s" then returnm FSGNJN_S
+ else if generic_eq p0_ "fsgnjx.s" then returnm FSGNJX_S
+ else if generic_eq p0_ "fmin.s" then returnm FMIN_S
+ else if generic_eq p0_ "fmax.s" then returnm FMAX_S
+ else if generic_eq p0_ "feq.s" then returnm FEQ_S
+ else if generic_eq p0_ "flt.s" then returnm FLT_S
+ else if generic_eq p0_ "fle.s" then returnm FLE_S
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_bin_op_S).
+
+Definition f_bin_type_mnemonic_S_forwards_matches (arg_ : f_bin_op_S) : bool :=
+ match arg_ with
+ | FSGNJ_S => true
+ | FSGNJN_S => true
+ | FSGNJX_S => true
+ | FMIN_S => true
+ | FMAX_S => true
+ | FEQ_S => true
+ | FLT_S => true
+ | FLE_S => true
+ end.
+
+Definition f_bin_type_mnemonic_S_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fsgnj.s" then true
+ else if generic_eq p0_ "fsgnjn.s" then true
+ else if generic_eq p0_ "fsgnjx.s" then true
+ else if generic_eq p0_ "fmin.s" then true
+ else if generic_eq p0_ "fmax.s" then true
+ else if generic_eq p0_ "feq.s" then true
+ else if generic_eq p0_ "flt.s" then true
+ else if generic_eq p0_ "fle.s" then true
+ else false.
+
+Definition _s1194_ (_s1195_ : string) : option string :=
+ let _s1196_ := _s1195_ in
+ if string_startswith _s1196_ "fle.s" then
+ match (string_drop _s1196_ (projT1 (string_length "fle.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1190_ (_s1191_ : string) : option string :=
+ let _s1192_ := _s1191_ in
+ if string_startswith _s1192_ "flt.s" then
+ match (string_drop _s1192_ (projT1 (string_length "flt.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1186_ (_s1187_ : string) : option string :=
+ let _s1188_ := _s1187_ in
+ if string_startswith _s1188_ "feq.s" then
+ match (string_drop _s1188_ (projT1 (string_length "feq.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1182_ (_s1183_ : string) : option string :=
+ let _s1184_ := _s1183_ in
+ if string_startswith _s1184_ "fmax.s" then
+ match (string_drop _s1184_ (projT1 (string_length "fmax.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1178_ (_s1179_ : string) : option string :=
+ let _s1180_ := _s1179_ in
+ if string_startswith _s1180_ "fmin.s" then
+ match (string_drop _s1180_ (projT1 (string_length "fmin.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1174_ (_s1175_ : string) : option string :=
+ let _s1176_ := _s1175_ in
+ if string_startswith _s1176_ "fsgnjx.s" then
+ match (string_drop _s1176_ (projT1 (string_length "fsgnjx.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1170_ (_s1171_ : string) : option string :=
+ let _s1172_ := _s1171_ in
+ if string_startswith _s1172_ "fsgnjn.s" then
+ match (string_drop _s1172_ (projT1 (string_length "fsgnjn.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1166_ (_s1167_ : string) : option string :=
+ let _s1168_ := _s1167_ in
+ if string_startswith _s1168_ "fsgnj.s" then
+ match (string_drop _s1168_ (projT1 (string_length "fsgnj.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_bin_type_mnemonic_S_matches_prefix (arg_ : string)
+: M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1169_ := arg_ in
+ (if match (_s1166_ _s1169_) with | Some s_ => true | _ => false end then
+ (match (_s1166_ _s1169_) with
+ | Some s_ =>
+ returnm (Some
+ (FSGNJ_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1170_ _s1169_) with | Some s_ => true | _ => false end then
+ (match (_s1170_ _s1169_) with
+ | Some s_ =>
+ returnm (Some
+ (FSGNJN_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1174_ _s1169_) with | Some s_ => true | _ => false end then
+ (match (_s1174_ _s1169_) with
+ | Some s_ =>
+ returnm (Some
+ (FSGNJX_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1178_ _s1169_) with | Some s_ => true | _ => false end then
+ (match (_s1178_ _s1169_) with
+ | Some s_ =>
+ returnm (Some
+ (FMIN_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1182_ _s1169_) with | Some s_ => true | _ => false end then
+ (match (_s1182_ _s1169_) with
+ | Some s_ =>
+ returnm (Some
+ (FMAX_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1186_ _s1169_) with | Some s_ => true | _ => false end then
+ (match (_s1186_ _s1169_) with
+ | Some s_ =>
+ returnm (Some
+ (FEQ_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1190_ _s1169_) with | Some s_ => true | _ => false end then
+ (match (_s1190_ _s1169_) with
+ | Some s_ =>
+ returnm (Some
+ (FLT_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1194_ _s1169_) with | Some s_ => true | _ => false end then
+ (match (_s1194_ _s1169_) with
+ | Some s_ =>
+ returnm (Some
+ (FLE_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_bin_op_S * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition f_un_type_mnemonic_S_forwards (arg_ : f_un_op_S) : string :=
+ match arg_ with | FMV_X_W => "fmv.x.w" | FCLASS_S => "fclass.s" | FMV_W_X => "fmv.w.x" end.
+
+Definition f_un_type_mnemonic_S_backwards (arg_ : string) : M (f_un_op_S) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fmv.x.w" then returnm FMV_X_W
+ else if generic_eq p0_ "fclass.s" then returnm FCLASS_S
+ else if generic_eq p0_ "fmv.w.x" then returnm FMV_W_X
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_un_op_S).
+
+Definition f_un_type_mnemonic_S_forwards_matches (arg_ : f_un_op_S) : bool :=
+ match arg_ with | FMV_X_W => true | FCLASS_S => true | FMV_W_X => true end.
+
+Definition f_un_type_mnemonic_S_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fmv.x.w" then true
+ else if generic_eq p0_ "fclass.s" then true
+ else if generic_eq p0_ "fmv.w.x" then true
+ else false.
+
+Definition _s1206_ (_s1207_ : string) : option string :=
+ let _s1208_ := _s1207_ in
+ if string_startswith _s1208_ "fmv.w.x" then
+ match (string_drop _s1208_ (projT1 (string_length "fmv.w.x"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1202_ (_s1203_ : string) : option string :=
+ let _s1204_ := _s1203_ in
+ if string_startswith _s1204_ "fclass.s" then
+ match (string_drop _s1204_ (projT1 (string_length "fclass.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1198_ (_s1199_ : string) : option string :=
+ let _s1200_ := _s1199_ in
+ if string_startswith _s1200_ "fmv.x.w" then
+ match (string_drop _s1200_ (projT1 (string_length "fmv.x.w"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_un_type_mnemonic_S_matches_prefix (arg_ : string)
+: M (option ((f_un_op_S * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1201_ := arg_ in
+ (if match (_s1198_ _s1201_) with | Some s_ => true | _ => false end then
+ (match (_s1198_ _s1201_) with
+ | Some s_ =>
+ returnm (Some
+ (FMV_X_W, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1202_ _s1201_) with | Some s_ => true | _ => false end then
+ (match (_s1202_ _s1201_) with
+ | Some s_ =>
+ returnm (Some
+ (FCLASS_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1206_ _s1201_) with | Some s_ => true | _ => false end then
+ (match (_s1206_ _s1201_) with
+ | Some s_ =>
+ returnm (Some
+ (FMV_W_X, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_op_S * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_op_S * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_un_op_S * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition fsplit_D (x64 : mword 64) : (mword 1 * mword 11 * mword 52) :=
+ (subrange_vec_dec x64 63 63, subrange_vec_dec x64 62 52, subrange_vec_dec x64 51 0).
+
+Definition fmake_D (sign : mword 1) (exp : mword 11) (mant : mword 52) : mword 64 :=
+ concat_vec sign (concat_vec exp mant).
+
+Definition canonical_NaN_D '(tt : unit) : mword 64 := Ox"7FF8000000000000" : mword 64.
+
+Definition f_is_neg_inf_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec sign ('b"1" : mword 1))
+ (andb (eq_vec exp (ones 11)) (eq_vec mant (zeros_implicit 52))).
+
+Definition f_is_neg_norm_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec sign ('b"1" : mword 1))
+ (andb (neq_vec exp (zeros_implicit 11)) (neq_vec exp (ones 11))).
+
+Definition f_is_neg_subnorm_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec sign ('b"1" : mword 1))
+ (andb (eq_vec exp (zeros_implicit 11)) (neq_vec mant (zeros_implicit 52))).
+
+Definition f_is_neg_zero_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec sign (ones 1))
+ (andb (eq_vec exp (zeros_implicit 11)) (eq_vec mant (zeros_implicit 52))).
+
+Definition f_is_pos_zero_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec sign (zeros_implicit 1))
+ (andb (eq_vec exp (zeros_implicit 11)) (eq_vec mant (zeros_implicit 52))).
+
+Definition f_is_pos_subnorm_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec sign (zeros_implicit 1))
+ (andb (eq_vec exp (zeros_implicit 11)) (neq_vec mant (zeros_implicit 52))).
+
+Definition f_is_pos_norm_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec sign (zeros_implicit 1))
+ (andb (neq_vec exp (zeros_implicit 11)) (neq_vec exp (ones 11))).
+
+Definition f_is_pos_inf_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec sign (zeros_implicit 1))
+ (andb (eq_vec exp (ones 11)) (eq_vec mant (zeros_implicit 52))).
+
+Definition f_is_SNaN_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec exp (ones 11))
+ (andb (eq_bit (access_vec_dec mant 51) B0) (neq_vec mant (zeros_implicit 52))).
+
+Definition f_is_QNaN_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec exp (ones 11)) (eq_bit (access_vec_dec mant 51) B1).
+
+Definition f_is_NaN_D (x64 : mword 64) : bool :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ andb (eq_vec exp (ones 11)) (neq_vec mant (zeros_implicit 52)).
+
+Definition negate_D (x64 : mword 64) : mword 64 :=
+ let '(sign, exp, mant) := fsplit_D x64 in
+ let new_sign := if eq_vec sign ('b"0" : mword 1) then 'b"1" : mword 1 else 'b"0" : mword 1 in
+ fmake_D new_sign exp mant.
+
+Definition feq_quiet_D (v1 : mword 64) (v2 : mword 64) : (bool * mword 5) :=
+ let '(s1, e1, m1) := fsplit_D v1 in
+ let '(s2, e2, m2) := fsplit_D v2 in
+ let v1Is0 := orb (f_is_neg_zero_D v1) (f_is_pos_zero_D v1) in
+ let v2Is0 := orb (f_is_neg_zero_D v2) (f_is_pos_zero_D v2) in
+ let result := orb (eq_vec v1 v2) (andb v1Is0 v2Is0) in
+ let fflags := if orb (f_is_SNaN_D v1) (f_is_SNaN_D v2) then nvFlag tt else zeros_implicit 5 in
+ (result, fflags).
+
+Definition flt_D (v1 : mword 64) (v2 : mword 64) (is_quiet : bool) : (bool * mword 5) :=
+ let '(s1, e1, m1) := fsplit_D v1 in
+ let '(s2, e2, m2) := fsplit_D v2 in
+ let result : bool :=
+ if andb (eq_vec s1 ('b"0" : mword 1)) (eq_vec s2 ('b"0" : mword 1)) then
+ if eq_vec e1 e2 then Z.ltb (projT1 (uint m1)) (projT1 (uint m2))
+ else Z.ltb (projT1 (uint e1)) (projT1 (uint e2))
+ else if andb (eq_vec s1 ('b"0" : mword 1)) (eq_vec s2 ('b"1" : mword 1)) then false
+ else if andb (eq_vec s1 ('b"1" : mword 1)) (eq_vec s2 ('b"0" : mword 1)) then true
+ else if eq_vec e1 e2 then Z.gtb (projT1 (uint m1)) (projT1 (uint m2))
+ else Z.gtb (projT1 (uint e1)) (projT1 (uint e2)) in
+ let fflags :=
+ if sumbool_of_bool is_quiet then
+ if orb (f_is_SNaN_D v1) (f_is_SNaN_D v2) then nvFlag tt
+ else zeros_implicit 5
+ else if orb (f_is_NaN_D v1) (f_is_NaN_D v2) then nvFlag tt
+ else zeros_implicit 5 in
+ (result, fflags).
+
+Definition fle_D (v1 : mword 64) (v2 : mword 64) (is_quiet : bool) : (bool * mword 5) :=
+ let '(s1, e1, m1) := fsplit_D v1 in
+ let '(s2, e2, m2) := fsplit_D v2 in
+ let v1Is0 := orb (f_is_neg_zero_D v1) (f_is_pos_zero_D v1) in
+ let v2Is0 := orb (f_is_neg_zero_D v2) (f_is_pos_zero_D v2) in
+ let result : bool :=
+ if andb (eq_vec s1 ('b"0" : mword 1)) (eq_vec s2 ('b"0" : mword 1)) then
+ if eq_vec e1 e2 then Z.leb (projT1 (uint m1)) (projT1 (uint m2))
+ else Z.ltb (projT1 (uint e1)) (projT1 (uint e2))
+ else if andb (eq_vec s1 ('b"0" : mword 1)) (eq_vec s2 ('b"1" : mword 1)) then
+ andb v1Is0 v2Is0
+ else if andb (eq_vec s1 ('b"1" : mword 1)) (eq_vec s2 ('b"0" : mword 1)) then true
+ else if eq_vec e1 e2 then Z.geb (projT1 (uint m1)) (projT1 (uint m2))
+ else Z.gtb (projT1 (uint e1)) (projT1 (uint e2)) in
+ let fflags :=
+ if sumbool_of_bool is_quiet then
+ if orb (f_is_SNaN_D v1) (f_is_SNaN_D v2) then nvFlag tt
+ else zeros_implicit 5
+ else if orb (f_is_NaN_D v1) (f_is_NaN_D v2) then nvFlag tt
+ else zeros_implicit 5 in
+ (result, fflags).
+
+Definition f_madd_type_mnemonic_D_forwards (arg_ : f_madd_op_D) : string :=
+ match arg_ with
+ | FMADD_D => "fmadd.d"
+ | FMSUB_D => "fmsub.d"
+ | FNMSUB_D => "fnmsub.d"
+ | FNMADD_D => "fnmadd.d"
+ end.
+
+Definition f_madd_type_mnemonic_D_backwards (arg_ : string) : M (f_madd_op_D) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fmadd.d" then returnm FMADD_D
+ else if generic_eq p0_ "fmsub.d" then returnm FMSUB_D
+ else if generic_eq p0_ "fnmsub.d" then returnm FNMSUB_D
+ else if generic_eq p0_ "fnmadd.d" then returnm FNMADD_D
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_madd_op_D).
+
+Definition f_madd_type_mnemonic_D_forwards_matches (arg_ : f_madd_op_D) : bool :=
+ match arg_ with | FMADD_D => true | FMSUB_D => true | FNMSUB_D => true | FNMADD_D => true end.
+
+Definition f_madd_type_mnemonic_D_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fmadd.d" then true
+ else if generic_eq p0_ "fmsub.d" then true
+ else if generic_eq p0_ "fnmsub.d" then true
+ else if generic_eq p0_ "fnmadd.d" then true
+ else false.
+
+Definition _s1222_ (_s1223_ : string) : option string :=
+ let _s1224_ := _s1223_ in
+ if string_startswith _s1224_ "fnmadd.d" then
+ match (string_drop _s1224_ (projT1 (string_length "fnmadd.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1218_ (_s1219_ : string) : option string :=
+ let _s1220_ := _s1219_ in
+ if string_startswith _s1220_ "fnmsub.d" then
+ match (string_drop _s1220_ (projT1 (string_length "fnmsub.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1214_ (_s1215_ : string) : option string :=
+ let _s1216_ := _s1215_ in
+ if string_startswith _s1216_ "fmsub.d" then
+ match (string_drop _s1216_ (projT1 (string_length "fmsub.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1210_ (_s1211_ : string) : option string :=
+ let _s1212_ := _s1211_ in
+ if string_startswith _s1212_ "fmadd.d" then
+ match (string_drop _s1212_ (projT1 (string_length "fmadd.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_madd_type_mnemonic_D_matches_prefix (arg_ : string)
+: M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1213_ := arg_ in
+ (if match (_s1210_ _s1213_) with | Some s_ => true | _ => false end then
+ (match (_s1210_ _s1213_) with
+ | Some s_ =>
+ returnm (Some
+ (FMADD_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1214_ _s1213_) with | Some s_ => true | _ => false end then
+ (match (_s1214_ _s1213_) with
+ | Some s_ =>
+ returnm (Some
+ (FMSUB_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1218_ _s1213_) with | Some s_ => true | _ => false end then
+ (match (_s1218_ _s1213_) with
+ | Some s_ =>
+ returnm (Some
+ (FNMSUB_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1222_ _s1213_) with | Some s_ => true | _ => false end then
+ (match (_s1222_ _s1213_) with
+ | Some s_ =>
+ returnm (Some
+ (FNMADD_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_madd_op_D * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition f_bin_rm_type_mnemonic_D_forwards (arg_ : f_bin_rm_op_D) : string :=
+ match arg_ with
+ | FADD_D => "fadd.d"
+ | FSUB_D => "fsub.d"
+ | FMUL_D => "fmul.d"
+ | FDIV_D => "fdiv.d"
+ end.
+
+Definition f_bin_rm_type_mnemonic_D_backwards (arg_ : string) : M (f_bin_rm_op_D) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fadd.d" then returnm FADD_D
+ else if generic_eq p0_ "fsub.d" then returnm FSUB_D
+ else if generic_eq p0_ "fmul.d" then returnm FMUL_D
+ else if generic_eq p0_ "fdiv.d" then returnm FDIV_D
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_bin_rm_op_D).
+
+Definition f_bin_rm_type_mnemonic_D_forwards_matches (arg_ : f_bin_rm_op_D) : bool :=
+ match arg_ with | FADD_D => true | FSUB_D => true | FMUL_D => true | FDIV_D => true end.
+
+Definition f_bin_rm_type_mnemonic_D_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fadd.d" then true
+ else if generic_eq p0_ "fsub.d" then true
+ else if generic_eq p0_ "fmul.d" then true
+ else if generic_eq p0_ "fdiv.d" then true
+ else false.
+
+Definition _s1238_ (_s1239_ : string) : option string :=
+ let _s1240_ := _s1239_ in
+ if string_startswith _s1240_ "fdiv.d" then
+ match (string_drop _s1240_ (projT1 (string_length "fdiv.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1234_ (_s1235_ : string) : option string :=
+ let _s1236_ := _s1235_ in
+ if string_startswith _s1236_ "fmul.d" then
+ match (string_drop _s1236_ (projT1 (string_length "fmul.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1230_ (_s1231_ : string) : option string :=
+ let _s1232_ := _s1231_ in
+ if string_startswith _s1232_ "fsub.d" then
+ match (string_drop _s1232_ (projT1 (string_length "fsub.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1226_ (_s1227_ : string) : option string :=
+ let _s1228_ := _s1227_ in
+ if string_startswith _s1228_ "fadd.d" then
+ match (string_drop _s1228_ (projT1 (string_length "fadd.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_bin_rm_type_mnemonic_D_matches_prefix (arg_ : string)
+: M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1229_ := arg_ in
+ (if match (_s1226_ _s1229_) with | Some s_ => true | _ => false end then
+ (match (_s1226_ _s1229_) with
+ | Some s_ =>
+ returnm (Some
+ (FADD_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1230_ _s1229_) with | Some s_ => true | _ => false end then
+ (match (_s1230_ _s1229_) with
+ | Some s_ =>
+ returnm (Some
+ (FSUB_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1234_ _s1229_) with | Some s_ => true | _ => false end then
+ (match (_s1234_ _s1229_) with
+ | Some s_ =>
+ returnm (Some
+ (FMUL_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1238_ _s1229_) with | Some s_ => true | _ => false end then
+ (match (_s1238_ _s1229_) with
+ | Some s_ =>
+ returnm (Some
+ (FDIV_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition f_un_rm_type_mnemonic_D_forwards (arg_ : f_un_rm_op_D) : string :=
+ match arg_ with
+ | FSQRT_D => "fsqrt.d"
+ | FCVT_W_D => "fcvt.w.d"
+ | FCVT_WU_D => "fcvt.wu.d"
+ | FCVT_D_W => "fcvt.d.w"
+ | FCVT_D_WU => "fcvt.d.wu"
+ | FCVT_L_D => "fcvt.l.d"
+ | FCVT_LU_D => "fcvt.lu.d"
+ | FCVT_D_L => "fcvt.d.l"
+ | FCVT_D_LU => "fcvt.d.lu"
+ | FCVT_S_D => "fcvt.s.d"
+ | FCVT_D_S => "fcvt.d.s"
+ end.
+
+Definition f_un_rm_type_mnemonic_D_backwards (arg_ : string) : M (f_un_rm_op_D) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fsqrt.d" then returnm FSQRT_D
+ else if generic_eq p0_ "fcvt.w.d" then returnm FCVT_W_D
+ else if generic_eq p0_ "fcvt.wu.d" then returnm FCVT_WU_D
+ else if generic_eq p0_ "fcvt.d.w" then returnm FCVT_D_W
+ else if generic_eq p0_ "fcvt.d.wu" then returnm FCVT_D_WU
+ else if generic_eq p0_ "fcvt.l.d" then returnm FCVT_L_D
+ else if generic_eq p0_ "fcvt.lu.d" then returnm FCVT_LU_D
+ else if generic_eq p0_ "fcvt.d.l" then returnm FCVT_D_L
+ else if generic_eq p0_ "fcvt.d.lu" then returnm FCVT_D_LU
+ else if generic_eq p0_ "fcvt.s.d" then returnm FCVT_S_D
+ else if generic_eq p0_ "fcvt.d.s" then returnm FCVT_D_S
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_un_rm_op_D).
+
+Definition f_un_rm_type_mnemonic_D_forwards_matches (arg_ : f_un_rm_op_D) : bool :=
+ match arg_ with
+ | FSQRT_D => true
+ | FCVT_W_D => true
+ | FCVT_WU_D => true
+ | FCVT_D_W => true
+ | FCVT_D_WU => true
+ | FCVT_L_D => true
+ | FCVT_LU_D => true
+ | FCVT_D_L => true
+ | FCVT_D_LU => true
+ | FCVT_S_D => true
+ | FCVT_D_S => true
+ end.
+
+Definition f_un_rm_type_mnemonic_D_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fsqrt.d" then true
+ else if generic_eq p0_ "fcvt.w.d" then true
+ else if generic_eq p0_ "fcvt.wu.d" then true
+ else if generic_eq p0_ "fcvt.d.w" then true
+ else if generic_eq p0_ "fcvt.d.wu" then true
+ else if generic_eq p0_ "fcvt.l.d" then true
+ else if generic_eq p0_ "fcvt.lu.d" then true
+ else if generic_eq p0_ "fcvt.d.l" then true
+ else if generic_eq p0_ "fcvt.d.lu" then true
+ else if generic_eq p0_ "fcvt.s.d" then true
+ else if generic_eq p0_ "fcvt.d.s" then true
+ else false.
+
+Definition _s1282_ (_s1283_ : string) : option string :=
+ let _s1284_ := _s1283_ in
+ if string_startswith _s1284_ "fcvt.d.s" then
+ match (string_drop _s1284_ (projT1 (string_length "fcvt.d.s"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1278_ (_s1279_ : string) : option string :=
+ let _s1280_ := _s1279_ in
+ if string_startswith _s1280_ "fcvt.s.d" then
+ match (string_drop _s1280_ (projT1 (string_length "fcvt.s.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1274_ (_s1275_ : string) : option string :=
+ let _s1276_ := _s1275_ in
+ if string_startswith _s1276_ "fcvt.d.lu" then
+ match (string_drop _s1276_ (projT1 (string_length "fcvt.d.lu"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1270_ (_s1271_ : string) : option string :=
+ let _s1272_ := _s1271_ in
+ if string_startswith _s1272_ "fcvt.d.l" then
+ match (string_drop _s1272_ (projT1 (string_length "fcvt.d.l"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1266_ (_s1267_ : string) : option string :=
+ let _s1268_ := _s1267_ in
+ if string_startswith _s1268_ "fcvt.lu.d" then
+ match (string_drop _s1268_ (projT1 (string_length "fcvt.lu.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1262_ (_s1263_ : string) : option string :=
+ let _s1264_ := _s1263_ in
+ if string_startswith _s1264_ "fcvt.l.d" then
+ match (string_drop _s1264_ (projT1 (string_length "fcvt.l.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1258_ (_s1259_ : string) : option string :=
+ let _s1260_ := _s1259_ in
+ if string_startswith _s1260_ "fcvt.d.wu" then
+ match (string_drop _s1260_ (projT1 (string_length "fcvt.d.wu"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1254_ (_s1255_ : string) : option string :=
+ let _s1256_ := _s1255_ in
+ if string_startswith _s1256_ "fcvt.d.w" then
+ match (string_drop _s1256_ (projT1 (string_length "fcvt.d.w"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1250_ (_s1251_ : string) : option string :=
+ let _s1252_ := _s1251_ in
+ if string_startswith _s1252_ "fcvt.wu.d" then
+ match (string_drop _s1252_ (projT1 (string_length "fcvt.wu.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1246_ (_s1247_ : string) : option string :=
+ let _s1248_ := _s1247_ in
+ if string_startswith _s1248_ "fcvt.w.d" then
+ match (string_drop _s1248_ (projT1 (string_length "fcvt.w.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1242_ (_s1243_ : string) : option string :=
+ let _s1244_ := _s1243_ in
+ if string_startswith _s1244_ "fsqrt.d" then
+ match (string_drop _s1244_ (projT1 (string_length "fsqrt.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_un_rm_type_mnemonic_D_matches_prefix (arg_ : string)
+: M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1245_ := arg_ in
+ (if match (_s1242_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1242_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FSQRT_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1246_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1246_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_W_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1250_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1250_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_WU_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1254_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1254_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_D_W, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1258_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1258_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_D_WU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1262_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1262_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_L_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1266_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1266_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_LU_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1270_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1270_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_D_L, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1274_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1274_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_D_LU, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1278_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1278_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_S_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1282_ _s1245_) with | Some s_ => true | _ => false end then
+ (match (_s1282_ _s1245_) with
+ | Some s_ =>
+ returnm (Some
+ (FCVT_D_S, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_un_rm_op_D * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition f_bin_type_mnemonic_D_forwards (arg_ : f_bin_op_D) : string :=
+ match arg_ with
+ | FSGNJ_D => "fsgnj.d"
+ | FSGNJN_D => "fsgnjn.d"
+ | FSGNJX_D => "fsgnjx.d"
+ | FMIN_D => "fmin.d"
+ | FMAX_D => "fmax.d"
+ | FEQ_D => "feq.d"
+ | FLT_D => "flt.d"
+ | FLE_D => "fle.d"
+ end.
+
+Definition f_bin_type_mnemonic_D_backwards (arg_ : string) : M (f_bin_op_D) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fsgnj.d" then returnm FSGNJ_D
+ else if generic_eq p0_ "fsgnjn.d" then returnm FSGNJN_D
+ else if generic_eq p0_ "fsgnjx.d" then returnm FSGNJX_D
+ else if generic_eq p0_ "fmin.d" then returnm FMIN_D
+ else if generic_eq p0_ "fmax.d" then returnm FMAX_D
+ else if generic_eq p0_ "feq.d" then returnm FEQ_D
+ else if generic_eq p0_ "flt.d" then returnm FLT_D
+ else if generic_eq p0_ "fle.d" then returnm FLE_D
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_bin_op_D).
+
+Definition f_bin_type_mnemonic_D_forwards_matches (arg_ : f_bin_op_D) : bool :=
+ match arg_ with
+ | FSGNJ_D => true
+ | FSGNJN_D => true
+ | FSGNJX_D => true
+ | FMIN_D => true
+ | FMAX_D => true
+ | FEQ_D => true
+ | FLT_D => true
+ | FLE_D => true
+ end.
+
+Definition f_bin_type_mnemonic_D_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fsgnj.d" then true
+ else if generic_eq p0_ "fsgnjn.d" then true
+ else if generic_eq p0_ "fsgnjx.d" then true
+ else if generic_eq p0_ "fmin.d" then true
+ else if generic_eq p0_ "fmax.d" then true
+ else if generic_eq p0_ "feq.d" then true
+ else if generic_eq p0_ "flt.d" then true
+ else if generic_eq p0_ "fle.d" then true
+ else false.
+
+Definition _s1314_ (_s1315_ : string) : option string :=
+ let _s1316_ := _s1315_ in
+ if string_startswith _s1316_ "fle.d" then
+ match (string_drop _s1316_ (projT1 (string_length "fle.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1310_ (_s1311_ : string) : option string :=
+ let _s1312_ := _s1311_ in
+ if string_startswith _s1312_ "flt.d" then
+ match (string_drop _s1312_ (projT1 (string_length "flt.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1306_ (_s1307_ : string) : option string :=
+ let _s1308_ := _s1307_ in
+ if string_startswith _s1308_ "feq.d" then
+ match (string_drop _s1308_ (projT1 (string_length "feq.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1302_ (_s1303_ : string) : option string :=
+ let _s1304_ := _s1303_ in
+ if string_startswith _s1304_ "fmax.d" then
+ match (string_drop _s1304_ (projT1 (string_length "fmax.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1298_ (_s1299_ : string) : option string :=
+ let _s1300_ := _s1299_ in
+ if string_startswith _s1300_ "fmin.d" then
+ match (string_drop _s1300_ (projT1 (string_length "fmin.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1294_ (_s1295_ : string) : option string :=
+ let _s1296_ := _s1295_ in
+ if string_startswith _s1296_ "fsgnjx.d" then
+ match (string_drop _s1296_ (projT1 (string_length "fsgnjx.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1290_ (_s1291_ : string) : option string :=
+ let _s1292_ := _s1291_ in
+ if string_startswith _s1292_ "fsgnjn.d" then
+ match (string_drop _s1292_ (projT1 (string_length "fsgnjn.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1286_ (_s1287_ : string) : option string :=
+ let _s1288_ := _s1287_ in
+ if string_startswith _s1288_ "fsgnj.d" then
+ match (string_drop _s1288_ (projT1 (string_length "fsgnj.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_bin_type_mnemonic_D_matches_prefix (arg_ : string)
+: M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1289_ := arg_ in
+ (if match (_s1286_ _s1289_) with | Some s_ => true | _ => false end then
+ (match (_s1286_ _s1289_) with
+ | Some s_ =>
+ returnm (Some
+ (FSGNJ_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1290_ _s1289_) with | Some s_ => true | _ => false end then
+ (match (_s1290_ _s1289_) with
+ | Some s_ =>
+ returnm (Some
+ (FSGNJN_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1294_ _s1289_) with | Some s_ => true | _ => false end then
+ (match (_s1294_ _s1289_) with
+ | Some s_ =>
+ returnm (Some
+ (FSGNJX_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1298_ _s1289_) with | Some s_ => true | _ => false end then
+ (match (_s1298_ _s1289_) with
+ | Some s_ =>
+ returnm (Some
+ (FMIN_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1302_ _s1289_) with | Some s_ => true | _ => false end then
+ (match (_s1302_ _s1289_) with
+ | Some s_ =>
+ returnm (Some
+ (FMAX_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1306_ _s1289_) with | Some s_ => true | _ => false end then
+ (match (_s1306_ _s1289_) with
+ | Some s_ =>
+ returnm (Some
+ (FEQ_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1310_ _s1289_) with | Some s_ => true | _ => false end then
+ (match (_s1310_ _s1289_) with
+ | Some s_ =>
+ returnm (Some
+ (FLT_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1314_ _s1289_) with | Some s_ => true | _ => false end then
+ (match (_s1314_ _s1289_) with
+ | Some s_ =>
+ returnm (Some
+ (FLE_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_bin_op_D * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition f_un_type_mnemonic_D_forwards (arg_ : f_un_op_D) : string :=
+ match arg_ with | FMV_X_D => "fmv.x.d" | FCLASS_D => "fclass.d" | FMV_D_X => "fmv.d.x" end.
+
+Definition f_un_type_mnemonic_D_backwards (arg_ : string) : M (f_un_op_D) :=
+ let p0_ := arg_ in
+ (if generic_eq p0_ "fmv.x.d" then returnm FMV_X_D
+ else if generic_eq p0_ "fclass.d" then returnm FCLASS_D
+ else if generic_eq p0_ "fmv.d.x" then returnm FMV_D_X
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (f_un_op_D).
+
+Definition f_un_type_mnemonic_D_forwards_matches (arg_ : f_un_op_D) : bool :=
+ match arg_ with | FMV_X_D => true | FCLASS_D => true | FMV_D_X => true end.
+
+Definition f_un_type_mnemonic_D_backwards_matches (arg_ : string) : bool :=
+ let p0_ := arg_ in
+ if generic_eq p0_ "fmv.x.d" then true
+ else if generic_eq p0_ "fclass.d" then true
+ else if generic_eq p0_ "fmv.d.x" then true
+ else false.
+
+Definition _s1326_ (_s1327_ : string) : option string :=
+ let _s1328_ := _s1327_ in
+ if string_startswith _s1328_ "fmv.d.x" then
+ match (string_drop _s1328_ (projT1 (string_length "fmv.d.x"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1322_ (_s1323_ : string) : option string :=
+ let _s1324_ := _s1323_ in
+ if string_startswith _s1324_ "fclass.d" then
+ match (string_drop _s1324_ (projT1 (string_length "fclass.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s1318_ (_s1319_ : string) : option string :=
+ let _s1320_ := _s1319_ in
+ if string_startswith _s1320_ "fmv.x.d" then
+ match (string_drop _s1320_ (projT1 (string_length "fmv.x.d"))) with | s_ => Some s_ end
+ else None.
+
+Definition f_un_type_mnemonic_D_matches_prefix (arg_ : string)
+: M (option ((f_un_op_D * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s1321_ := arg_ in
+ (if match (_s1318_ _s1321_) with | Some s_ => true | _ => false end then
+ (match (_s1318_ _s1321_) with
+ | Some s_ =>
+ returnm (Some
+ (FMV_X_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1322_ _s1321_) with | Some s_ => true | _ => false end then
+ (match (_s1322_ _s1321_) with
+ | Some s_ =>
+ returnm (Some
+ (FCLASS_D, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_op_D * {n : Z & ArithFact (n >=? 0)})))
+ end)
+ : M (option ((f_un_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s1326_ _s1321_) with | Some s_ => true | _ => false end then
+ (match (_s1326_ _s1321_) with
+ | Some s_ =>
+ returnm (Some
+ (FMV_D_X, build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((f_un_op_D * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((csrop * {n : Z & ArithFact (n >= 0)})))
- else returnm (None : option ((csrop * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((csrop * {n : Z & ArithFact (n >= 0)}))).
+ : M (option ((f_un_op_D * {n : Z & ArithFact (n >=? 0)})))
+ else returnm None)
+ : M (option ((f_un_op_D * {n : Z & ArithFact (n >=? 0)}))).
-Definition encdec_forwards (arg_ : ast)
-: M (mword 32) :=
-
+Definition encdec_forwards (arg_ : ast) : M (mword 32) :=
(match arg_ with
- | UTYPE ((imm, rd, op)) =>
- returnm ((concat_vec (imm : mword 20) (concat_vec (rd : mword 5) (encdec_uop_forwards op)))
- : mword (20 + (5 + 7)))
- | RISCV_JAL ((v__2, rd)) =>
- (if ((eq_vec (subrange_vec_dec v__2 0 0) (vec_of_bits [B0] : mword (0 - 0 + 1)))) then
+ | UTYPE (imm, rd, op) =>
+ returnm (concat_vec (imm : mword 20) (concat_vec (rd : mword 5) (encdec_uop_forwards op)))
+ | RISCV_JAL (v__2, rd) =>
+ (if eq_vec (subrange_vec_dec v__2 0 0) ('b"0" : mword (0 - 0 + 1)) then
let imm_19 : bits 1 := subrange_vec_dec v__2 20 20 in
let imm_8 : bits 1 := subrange_vec_dec v__2 11 11 in
let imm_7_0 : bits 8 := subrange_vec_dec v__2 19 12 in
let imm_19 : bits 1 := subrange_vec_dec v__2 20 20 in
let imm_18_13 : bits 6 := subrange_vec_dec v__2 10 5 in
let imm_12_9 : bits 4 := subrange_vec_dec v__2 4 1 in
- returnm ((concat_vec (imm_19 : bits 1)
- (concat_vec (imm_18_13 : bits 6)
- (concat_vec (imm_12_9 : bits 4)
- (concat_vec (imm_8 : bits 1)
- (concat_vec (imm_7_0 : bits 8)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B1;B1;B0;B1;B1;B1;B1] : mword 7)))))))
- : mword 32)
+ returnm (concat_vec (imm_19 : bits 1)
+ (concat_vec (imm_18_13 : bits 6)
+ (concat_vec (imm_12_9 : bits 4)
+ (concat_vec (imm_8 : bits 1)
+ (concat_vec (imm_7_0 : bits 8)
+ (concat_vec (rd : mword 5) ('b"1101111" : mword 7)))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | RISCV_JALR ((imm, rs1, rd)) =>
- returnm ((concat_vec (imm : mword 12)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (rd : mword 5) (vec_of_bits [B1;B1;B0;B0;B1;B1;B1] : mword 7)))))
- : mword (12 + (5 + (3 + (5 + 7)))))
- | BTYPE ((v__4, rs2, rs1, op)) =>
- (if ((eq_vec (subrange_vec_dec v__4 0 0) (vec_of_bits [B0] : mword (0 - 0 + 1)))) then
+ | RISCV_JALR (imm, rs1, rd) =>
+ returnm (concat_vec (imm : mword 12)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1100111" : mword 7)))))
+ | BTYPE (v__4, rs2, rs1, op) =>
+ (if eq_vec (subrange_vec_dec v__4 0 0) ('b"0" : mword (0 - 0 + 1)) then
let imm7_6 : bits 1 := subrange_vec_dec v__4 12 12 in
let imm7_6 : bits 1 := subrange_vec_dec v__4 12 12 in
let imm7_5_0 : bits 6 := subrange_vec_dec v__4 10 5 in
let imm5_4_1 : bits 4 := subrange_vec_dec v__4 4 1 in
let imm5_0 : bits 1 := subrange_vec_dec v__4 11 11 in
- returnm ((concat_vec (imm7_6 : bits 1)
- (concat_vec (imm7_5_0 : bits 6)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (encdec_bop_forwards op)
- (concat_vec (imm5_4_1 : bits 4)
- (concat_vec (imm5_0 : bits 1)
- (vec_of_bits [B1;B1;B0;B0;B0;B1;B1] : mword 7))))))))
- : mword 32)
+ returnm (concat_vec (imm7_6 : bits 1)
+ (concat_vec (imm7_5_0 : bits 6)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_bop_forwards op)
+ (concat_vec (imm5_4_1 : bits 4)
+ (concat_vec (imm5_0 : bits 1) ('b"1100011" : mword 7))))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | ITYPE ((imm, rs1, rd, op)) =>
- returnm ((concat_vec (imm : mword 12)
- (concat_vec (rs1 : mword 5)
- (concat_vec (encdec_iop_forwards op)
- (concat_vec (rd : mword 5) (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7)))))
- : mword (12 + (5 + (3 + (5 + 7)))))
- | SHIFTIOP ((shamt, rs1, rd, RISCV_SLLI)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)
- (concat_vec (shamt : mword 6)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (6 + (6 + (5 + (3 + (5 + 7))))))
- | SHIFTIOP ((shamt, rs1, rd, RISCV_SRLI)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)
- (concat_vec (shamt : mword 6)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (6 + (6 + (5 + (3 + (5 + 7))))))
- | SHIFTIOP ((shamt, rs1, rd, RISCV_SRAI)) =>
- returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0] : mword 6)
- (concat_vec (shamt : mword 6)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (6 + (6 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_ADD)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_SLT)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B1;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_SLTU)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B1;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_AND)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B1;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_OR)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B1;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_XOR)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_SLL)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_SRL)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_SUB)) =>
- returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | RTYPE ((rs2, rs1, rd, RISCV_SRA)) =>
- returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | LOAD ((imm, rs1, rd, is_unsigned, size, false, false)) =>
- (if sumbool_of_bool ((orb (Z.ltb (projT1 (word_width_bytes size)) 8)
- (andb (negb is_unsigned)
- ((Z.leb (projT1 (word_width_bytes size)) 8)
- : bool)))) then
- returnm ((concat_vec (imm : mword 12)
- (concat_vec (rs1 : mword 5)
- (concat_vec (bool_bits_forwards is_unsigned)
- (concat_vec (size_bits_forwards size)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B0;B0;B0;B1;B1] : mword 7))))))
- : mword (12 + (5 + (1 + (2 + (5 + 7))))))
+ | ITYPE (imm, rs1, rd, op) =>
+ returnm (concat_vec (imm : mword 12)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_iop_forwards op)
+ (concat_vec (rd : mword 5) ('b"0010011" : mword 7)))))
+ | SHIFTIOP (shamt, rs1, rd, RISCV_SLLI) =>
+ returnm (concat_vec ('b"000000" : mword 6)
+ (concat_vec (shamt : mword 6)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0010011" : mword 7))))))
+ | SHIFTIOP (shamt, rs1, rd, RISCV_SRLI) =>
+ returnm (concat_vec ('b"000000" : mword 6)
+ (concat_vec (shamt : mword 6)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0010011" : mword 7))))))
+ | SHIFTIOP (shamt, rs1, rd, RISCV_SRAI) =>
+ returnm (concat_vec ('b"010000" : mword 6)
+ (concat_vec (shamt : mword 6)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0010011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_ADD) =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_SLT) =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"010" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_SLTU) =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"011" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_AND) =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"111" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_OR) =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"110" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_XOR) =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"100" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_SLL) =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_SRL) =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_SUB) =>
+ returnm (concat_vec ('b"0100000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | RTYPE (rs2, rs1, rd, RISCV_SRA) =>
+ returnm (concat_vec ('b"0100000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | LOAD (imm, rs1, rd, is_unsigned, size, false, false) =>
+ (if sumbool_of_bool
+ (orb (Z.ltb (projT1 (word_width_bytes size)) 8)
+ (andb (negb is_unsigned) ((Z.leb (projT1 (word_width_bytes size)) 8) : bool))) then
+ returnm (concat_vec (imm : mword 12)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (bool_bits_forwards is_unsigned)
+ (concat_vec (size_bits_forwards size)
+ (concat_vec (rd : mword 5) ('b"0000011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | STORE ((v__6, rs2, rs1, size, false, false)) =>
- (if sumbool_of_bool ((Z.leb (projT1 (word_width_bytes size)) 8)) then
+ | STORE (v__6, rs2, rs1, size, false, false) =>
+ (if sumbool_of_bool (Z.leb (projT1 (word_width_bytes size)) 8) then
let imm7 : bits 7 := subrange_vec_dec v__6 11 5 in
let imm7 : bits 7 := subrange_vec_dec v__6 11 5 in
let imm5 : bits 5 := subrange_vec_dec v__6 4 0 in
- returnm ((concat_vec (imm7 : bits 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (size_bits_forwards size)
- (concat_vec (imm5 : bits 5)
- (vec_of_bits [B0;B1;B0;B0;B0;B1;B1] : mword 7)))))))
- : mword 32)
+ returnm (concat_vec (imm7 : bits 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec (size_bits_forwards size)
+ (concat_vec (imm5 : bits 5) ('b"0100011" : mword 7)))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | ADDIW ((imm, rs1, rd)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (imm : mword 12)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7)))))
- : mword (12 + (5 + (3 + (5 + 7)))))
+ | ADDIW (imm, rs1, rd) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec (imm : mword 12)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0011011" : mword 7)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | SHIFTW ((shamt, rs1, rd, RISCV_SLLI)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (shamt : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | SHIFTW (shamt, rs1, rd, RISCV_SLLI) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (shamt : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0011011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | SHIFTW ((shamt, rs1, rd, RISCV_SRLI)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (shamt : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | SHIFTW (shamt, rs1, rd, RISCV_SRLI) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (shamt : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0011011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | SHIFTW ((shamt, rs1, rd, RISCV_SRAI)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (shamt : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | SHIFTW (shamt, rs1, rd, RISCV_SRAI) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0100000" : mword 7)
+ (concat_vec (shamt : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0011011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | RTYPEW ((rs2, rs1, rd, RISCV_ADDW)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | RTYPEW (rs2, rs1, rd, RISCV_ADDW) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0111011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | RTYPEW ((rs2, rs1, rd, RISCV_SUBW)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | RTYPEW (rs2, rs1, rd, RISCV_SUBW) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0100000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0111011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | RTYPEW ((rs2, rs1, rd, RISCV_SLLW)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | RTYPEW (rs2, rs1, rd, RISCV_SLLW) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0111011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | RTYPEW ((rs2, rs1, rd, RISCV_SRLW)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | RTYPEW (rs2, rs1, rd, RISCV_SRLW) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0111011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | RTYPEW ((rs2, rs1, rd, RISCV_SRAW)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | RTYPEW (rs2, rs1, rd, RISCV_SRAW) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0100000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0111011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | SHIFTIWOP ((shamt, rs1, rd, RISCV_SLLIW)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (shamt : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | SHIFTIWOP (shamt, rs1, rd, RISCV_SLLIW) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (shamt : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0011011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | SHIFTIWOP ((shamt, rs1, rd, RISCV_SRLIW)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (shamt : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | SHIFTIWOP (shamt, rs1, rd, RISCV_SRLIW) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (shamt : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0011011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | SHIFTIWOP ((shamt, rs1, rd, RISCV_SRAIW)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (shamt : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ | SHIFTIWOP (shamt, rs1, rd, RISCV_SRAIW) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0100000" : mword 7)
+ (concat_vec (shamt : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"101" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0011011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | FENCE ((pred, succ)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0] : mword 4)
- (concat_vec (pred : mword 4)
- (concat_vec (succ : mword 4)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7)))))))
- : mword (4 + (4 + (4 + (5 + (3 + (5 + 7)))))))
- | FENCE_TSO ((pred, succ)) =>
- returnm ((concat_vec (vec_of_bits [B1;B0;B0;B0] : mword 4)
- (concat_vec (pred : mword 4)
- (concat_vec (succ : mword 4)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7)))))))
- : mword (4 + (4 + (4 + (5 + (3 + (5 + 7)))))))
- | FENCEI (tt) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B0;B0;B0;B1;B1;B1;B1] : mword 7)))))
- : mword (12 + (5 + (3 + (5 + 7)))))
- | ECALL (tt) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7)))))
- : mword (12 + (5 + (3 + (5 + 7)))))
- | MRET (tt) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B1;B1;B0;B0;B0] : mword 7)
- (concat_vec (vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | SRET (tt) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B0;B0;B0] : mword 7)
- (concat_vec (vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | EBREAK (tt) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 12)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7)))))
- : mword (12 + (5 + (3 + (5 + 7)))))
- | WFI (tt) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1] : mword 12)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7)))))
- : mword (12 + (5 + (3 + (5 + 7)))))
- | SFENCE_VMA ((rs1, rs2)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B0;B0;B1] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | LOADRES ((aq, rl, rs1, size, rd)) =>
- (if sumbool_of_bool ((Z.leb (projT1 (word_width_bytes size)) 8)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
- (concat_vec (bool_bits_forwards aq)
- (concat_vec (bool_bits_forwards rl)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (size_bits_forwards size)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))))))
- : mword (5 + (1 + (1 + (5 + (5 + (1 + (2 + (5 + 7)))))))))
+ | FENCE (pred, succ) =>
+ returnm (concat_vec (Ox"0" : mword 4)
+ (concat_vec (pred : mword 4)
+ (concat_vec (succ : mword 4)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"0001111" : mword 7)))))))
+ | FENCE_TSO (pred, succ) =>
+ returnm (concat_vec (Ox"8" : mword 4)
+ (concat_vec (pred : mword 4)
+ (concat_vec (succ : mword 4)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"0001111" : mword 7)))))))
+ | FENCEI tt =>
+ returnm (concat_vec (Ox"000" : mword 12)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"0001111" : mword 7)))))
+ | ECALL tt =>
+ returnm (concat_vec (Ox"000" : mword 12)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"1110011" : mword 7)))))
+ | MRET tt =>
+ returnm (concat_vec ('b"0011000" : mword 7)
+ (concat_vec ('b"00010" : mword 5)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"1110011" : mword 7))))))
+ | SRET tt =>
+ returnm (concat_vec ('b"0001000" : mword 7)
+ (concat_vec ('b"00010" : mword 5)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"1110011" : mword 7))))))
+ | EBREAK tt =>
+ returnm (concat_vec (Ox"001" : mword 12)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"1110011" : mword 7)))))
+ | WFI tt =>
+ returnm (concat_vec (Ox"105" : mword 12)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"1110011" : mword 7)))))
+ | SFENCE_VMA (rs1, rs2) =>
+ returnm (concat_vec ('b"0001001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"1110011" : mword 7))))))
+ | LOADRES (aq, rl, rs1, size, rd) =>
+ (if sumbool_of_bool (Z.leb (projT1 (word_width_bytes size)) 8) then
+ returnm (concat_vec ('b"00010" : mword 5)
+ (concat_vec (bool_bits_forwards aq)
+ (concat_vec (bool_bits_forwards rl)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec (size_bits_forwards size)
+ (concat_vec (rd : mword 5) ('b"0101111" : mword 7)))))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | STORECON ((aq, rl, rs2, rs1, size, rd)) =>
- (if sumbool_of_bool ((Z.leb (projT1 (word_width_bytes size)) 8)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B1;B1] : mword 5)
- (concat_vec (bool_bits_forwards aq)
- (concat_vec (bool_bits_forwards rl)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (size_bits_forwards size)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))))))
- : mword (5 + (1 + (1 + (5 + (5 + (1 + (2 + (5 + 7)))))))))
+ | STORECON (aq, rl, rs2, rs1, size, rd) =>
+ (if sumbool_of_bool (Z.leb (projT1 (word_width_bytes size)) 8) then
+ returnm (concat_vec ('b"00011" : mword 5)
+ (concat_vec (bool_bits_forwards aq)
+ (concat_vec (bool_bits_forwards rl)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec (size_bits_forwards size)
+ (concat_vec (rd : mword 5) ('b"0101111" : mword 7)))))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | AMO ((op, aq, rl, rs2, rs1, size, rd)) =>
- (if sumbool_of_bool ((Z.leb (projT1 (word_width_bytes size)) 8)) then
- returnm ((concat_vec (encdec_amoop_forwards op)
- (concat_vec (bool_bits_forwards aq)
- (concat_vec (bool_bits_forwards rl)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (size_bits_forwards size)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1] : mword 7)))))))))
- : mword (5 + (1 + (1 + (5 + (5 + (1 + (2 + (5 + 7)))))))))
+ | AMO (op, aq, rl, rs2, rs1, size, rd) =>
+ (if sumbool_of_bool (Z.leb (projT1 (word_width_bytes size)) 8) then
+ returnm (concat_vec (encdec_amoop_forwards op)
+ (concat_vec (bool_bits_forwards aq)
+ (concat_vec (bool_bits_forwards rl)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec (size_bits_forwards size)
+ (concat_vec (rd : mword 5) ('b"0101111" : mword 7)))))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | MUL ((rs2, rs1, rd, high, signed1, signed2)) =>
+ | MUL (rs2, rs1, rd, high, signed1, signed2) =>
(encdec_mul_op_forwards (high, signed1, signed2)) >>= fun w__38 : mword 3 =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (w__38 : bits 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | DIV ((rs2, rs1, rd, s)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0] : mword 2)
- (concat_vec (bool_not_bits_forwards s)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))))))
- : mword (7 + (5 + (5 + (2 + (1 + (5 + 7)))))))
- | REM ((rs2, rs1, rd, s)) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (bool_not_bits_forwards s)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword 7)))))))
- : mword (7 + (5 + (5 + (2 + (1 + (5 + 7)))))))
- | MULW ((rs2, rs1, rd)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
+ returnm (concat_vec ('b"0000001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (w__38 : bits 3)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7))))))
+ | DIV (rs2, rs1, rd, s) =>
+ returnm (concat_vec ('b"0000001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"10" : mword 2)
+ (concat_vec (bool_not_bits_forwards s)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7)))))))
+ | REM (rs2, rs1, rd, s) =>
+ returnm (concat_vec ('b"0000001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (bool_not_bits_forwards s)
+ (concat_vec (rd : mword 5) ('b"0110011" : mword 7)))))))
+ | MULW (rs2, rs1, rd) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0111011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | DIVW (rs2, rs1, rd, s) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"10" : mword 2)
+ (concat_vec (bool_not_bits_forwards s)
+ (concat_vec (rd : mword 5) ('b"0111011" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | REMW (rs2, rs1, rd, s) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"0000001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (bool_not_bits_forwards s)
+ (concat_vec (rd : mword 5) ('b"0111011" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | CSR (csr, rs1, rd, is_imm, op) =>
+ returnm (concat_vec (csr : mword 12)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (bool_bits_forwards is_imm)
+ (concat_vec (encdec_csrop_forwards op)
+ (concat_vec (rd : mword 5) ('b"1110011" : mword 7))))))
+ | URET tt =>
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec ('b"00010" : mword 5)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"00000" : mword 5) ('b"1110011" : mword 7))))))
+ | LOAD_FP (imm, rs1, rd, WORD) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__45 : bool =>
+ (if sumbool_of_bool w__45 then
+ returnm (concat_vec (imm : mword 12)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"010" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0000111" : mword 7)))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | LOAD_FP (imm, rs1, rd, DOUBLE) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__48 : bool =>
+ (if sumbool_of_bool w__48 then
+ returnm (concat_vec (imm : mword 12)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"011" : mword 3)
+ (concat_vec (rd : mword 5) ('b"0000111" : mword 7)))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | STORE_FP (v__7, rs2, rs1, WORD) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__51 : bool =>
+ (if sumbool_of_bool w__51 then
+ let imm7 : bits 7 := subrange_vec_dec v__7 11 5 in
+ let imm7 : bits 7 := subrange_vec_dec v__7 11 5 in
+ let imm5 : bits 5 := subrange_vec_dec v__7 4 0 in
+ returnm (concat_vec (imm7 : bits 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"010" : mword 3)
+ (concat_vec (imm5 : bits 5) ('b"0100111" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | STORE_FP (v__8, rs2, rs1, DOUBLE) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__54 : bool =>
+ (if sumbool_of_bool w__54 then
+ let imm7 : bits 7 := subrange_vec_dec v__8 11 5 in
+ let imm7 : bits 7 := subrange_vec_dec v__8 11 5 in
+ let imm5 : bits 5 := subrange_vec_dec v__8 4 0 in
+ returnm (concat_vec (imm7 : bits 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"011" : mword 3)
+ (concat_vec (imm5 : bits 5) ('b"0100111" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FMADD_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__57 : bool =>
+ (if sumbool_of_bool w__57 then
+ returnm (concat_vec (rs3 : mword 5)
+ (concat_vec ('b"00" : mword 2)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1000011" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FMSUB_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__60 : bool =>
+ (if sumbool_of_bool w__60 then
+ returnm (concat_vec (rs3 : mword 5)
+ (concat_vec ('b"00" : mword 2)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1000111" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FNMSUB_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__63 : bool =>
+ (if sumbool_of_bool w__63 then
+ returnm (concat_vec (rs3 : mword 5)
+ (concat_vec ('b"00" : mword 2)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1001011" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FNMADD_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__66 : bool =>
+ (if sumbool_of_bool w__66 then
+ returnm (concat_vec (rs3 : mword 5)
+ (concat_vec ('b"00" : mword 2)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1001111" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FADD_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__69 : bool =>
+ (if sumbool_of_bool w__69 then
+ returnm (concat_vec ('b"0000000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FSUB_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__72 : bool =>
+ (if sumbool_of_bool w__72 then
+ returnm (concat_vec ('b"0000100" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FMUL_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__75 : bool =>
+ (if sumbool_of_bool w__75 then
+ returnm (concat_vec ('b"0001000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FDIV_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__78 : bool =>
+ (if sumbool_of_bool w__78 then
+ returnm (concat_vec ('b"0001100" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FSQRT_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__81 : bool =>
+ (if sumbool_of_bool w__81 then
+ returnm (concat_vec ('b"0101100" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_W_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__84 : bool =>
+ (if sumbool_of_bool w__84 then
+ returnm (concat_vec ('b"1100000" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_WU_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__87 : bool =>
+ (if sumbool_of_bool w__87 then
+ returnm (concat_vec ('b"1100000" : mword 7)
+ (concat_vec ('b"00001" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_W) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__90 : bool =>
+ (if sumbool_of_bool w__90 then
+ returnm (concat_vec ('b"1101000" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_WU) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__93 : bool =>
+ (if sumbool_of_bool w__93 then
+ returnm (concat_vec ('b"1101000" : mword 7)
+ (concat_vec ('b"00001" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_L_S) =>
+ (is_RV64F tt) >>= fun w__96 : bool =>
+ (if sumbool_of_bool w__96 then
+ returnm (concat_vec ('b"1100000" : mword 7)
+ (concat_vec ('b"00010" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_LU_S) =>
+ (is_RV64F tt) >>= fun w__99 : bool =>
+ (if sumbool_of_bool w__99 then
+ returnm (concat_vec ('b"1100000" : mword 7)
+ (concat_vec ('b"00011" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_L) =>
+ (is_RV64F tt) >>= fun w__102 : bool =>
+ (if sumbool_of_bool w__102 then
+ returnm (concat_vec ('b"1101000" : mword 7)
+ (concat_vec ('b"00010" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_LU) =>
+ (is_RV64F tt) >>= fun w__105 : bool =>
+ (if sumbool_of_bool w__105 then
+ returnm (concat_vec ('b"1101000" : mword 7)
+ (concat_vec ('b"00011" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJ_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__108 : bool =>
+ (if sumbool_of_bool w__108 then
+ returnm (concat_vec ('b"0010000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJN_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__111 : bool =>
+ (if sumbool_of_bool w__111 then
+ returnm (concat_vec ('b"0010000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJX_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__114 : bool =>
+ (if sumbool_of_bool w__114 then
+ returnm (concat_vec ('b"0010000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"010" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FMIN_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__117 : bool =>
+ (if sumbool_of_bool w__117 then
+ returnm (concat_vec ('b"0010100" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FMAX_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__120 : bool =>
+ (if sumbool_of_bool w__120 then
+ returnm (concat_vec ('b"0010100" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FEQ_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__123 : bool =>
+ (if sumbool_of_bool w__123 then
+ returnm (concat_vec ('b"1010000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"010" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FLT_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__126 : bool =>
+ (if sumbool_of_bool w__126 then
+ returnm (concat_vec ('b"1010000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FLE_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__129 : bool =>
+ (if sumbool_of_bool w__129 then
+ returnm (concat_vec ('b"1010000" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_TYPE_S (rs1, rd, FCLASS_S) =>
+ (haveFExt tt) >>= fun w__132 : bool =>
+ (if sumbool_of_bool w__132 then
+ returnm (concat_vec ('b"1110000" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_TYPE_S (rs1, rd, FMV_X_W) =>
+ (haveFExt tt) >>= fun w__135 : bool =>
+ (if sumbool_of_bool w__135 then
+ returnm (concat_vec ('b"1110000" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_TYPE_S (rs1, rd, FMV_W_X) =>
+ (haveFExt tt) >>= fun w__138 : bool =>
+ (if sumbool_of_bool w__138 then
+ returnm (concat_vec ('b"1111000" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, FMADD_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__141 : bool =>
+ (if sumbool_of_bool w__141 then
+ returnm (concat_vec (rs3 : mword 5)
+ (concat_vec ('b"01" : mword 2)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1000011" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, FMSUB_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__144 : bool =>
+ (if sumbool_of_bool w__144 then
+ returnm (concat_vec (rs3 : mword 5)
+ (concat_vec ('b"01" : mword 2)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1000111" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, FNMSUB_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__147 : bool =>
+ (if sumbool_of_bool w__147 then
+ returnm (concat_vec (rs3 : mword 5)
+ (concat_vec ('b"01" : mword 2)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1001011" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, FNMADD_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__150 : bool =>
+ (if sumbool_of_bool w__150 then
+ returnm (concat_vec (rs3 : mword 5)
+ (concat_vec ('b"01" : mword 2)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1001111" : mword 7)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, FADD_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__153 : bool =>
+ (if sumbool_of_bool w__153 then
+ returnm (concat_vec ('b"0000001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, FSUB_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__156 : bool =>
+ (if sumbool_of_bool w__156 then
+ returnm (concat_vec ('b"0000101" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, FMUL_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__159 : bool =>
+ (if sumbool_of_bool w__159 then
+ returnm (concat_vec ('b"0001001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, FDIV_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__162 : bool =>
+ (if sumbool_of_bool w__162 then
+ returnm (concat_vec ('b"0001101" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | DIVW ((rs2, rs1, rd, s)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B0] : mword 2)
- (concat_vec (bool_not_bits_forwards s)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))))))
- : mword (7 + (5 + (5 + (2 + (1 + (5 + 7)))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FSQRT_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__165 : bool =>
+ (if sumbool_of_bool w__165 then
+ returnm (concat_vec ('b"0101101" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | REMW ((rs2, rs1, rd, s)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B1] : mword 7)
- (concat_vec (rs2 : mword 5)
- (concat_vec (rs1 : mword 5)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (bool_not_bits_forwards s)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1] : mword 7)))))))
- : mword (7 + (5 + (5 + (2 + (1 + (5 + 7)))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_W_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__168 : bool =>
+ (if sumbool_of_bool w__168 then
+ returnm (concat_vec ('b"1100001" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 32)
- | CSR ((csr, rs1, rd, is_imm, op)) =>
- returnm ((concat_vec (csr : mword 12)
- (concat_vec (rs1 : mword 5)
- (concat_vec (bool_bits_forwards is_imm)
- (concat_vec (encdec_csrop_forwards op)
- (concat_vec (rd : mword 5)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (12 + (5 + (1 + (2 + (5 + 7))))))
- | URET (tt) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword 7)
- (concat_vec (vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1] : mword 7))))))
- : mword (7 + (5 + (5 + (3 + (5 + 7))))))
- | ILLEGAL (s) => returnm (s : mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_WU_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__171 : bool =>
+ (if sumbool_of_bool w__171 then
+ returnm (concat_vec ('b"1100001" : mword 7)
+ (concat_vec ('b"00001" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_W) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__174 : bool =>
+ (if sumbool_of_bool w__174 then
+ returnm (concat_vec ('b"1101001" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_WU) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__177 : bool =>
+ (if sumbool_of_bool w__177 then
+ returnm (concat_vec ('b"1101001" : mword 7)
+ (concat_vec ('b"00001" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_S_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__180 : bool =>
+ (if sumbool_of_bool w__180 then
+ returnm (concat_vec ('b"0100000" : mword 7)
+ (concat_vec ('b"00001" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_S) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__183 : bool =>
+ (if sumbool_of_bool w__183 then
+ returnm (concat_vec ('b"0100001" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_L_D) =>
+ (is_RV64D tt) >>= fun w__186 : bool =>
+ (if sumbool_of_bool w__186 then
+ returnm (concat_vec ('b"1100001" : mword 7)
+ (concat_vec ('b"00010" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_LU_D) =>
+ (is_RV64D tt) >>= fun w__189 : bool =>
+ (if sumbool_of_bool w__189 then
+ returnm (concat_vec ('b"1100001" : mword 7)
+ (concat_vec ('b"00011" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_L) =>
+ (is_RV64D tt) >>= fun w__192 : bool =>
+ (if sumbool_of_bool w__192 then
+ returnm (concat_vec ('b"1101001" : mword 7)
+ (concat_vec ('b"00010" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_LU) =>
+ (is_RV64D tt) >>= fun w__195 : bool =>
+ (if sumbool_of_bool w__195 then
+ returnm (concat_vec ('b"1101001" : mword 7)
+ (concat_vec ('b"00011" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec (encdec_rounding_mode_forwards rm)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJ_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__198 : bool =>
+ (if sumbool_of_bool w__198 then
+ returnm (concat_vec ('b"0010001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJN_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__201 : bool =>
+ (if sumbool_of_bool w__201 then
+ returnm (concat_vec ('b"0010001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJX_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__204 : bool =>
+ (if sumbool_of_bool w__204 then
+ returnm (concat_vec ('b"0010001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"010" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FMIN_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__207 : bool =>
+ (if sumbool_of_bool w__207 then
+ returnm (concat_vec ('b"0010101" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FMAX_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__210 : bool =>
+ (if sumbool_of_bool w__210 then
+ returnm (concat_vec ('b"0010101" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FEQ_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__213 : bool =>
+ (if sumbool_of_bool w__213 then
+ returnm (concat_vec ('b"1010001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"010" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FLT_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__216 : bool =>
+ (if sumbool_of_bool w__216 then
+ returnm (concat_vec ('b"1010001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FLE_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__219 : bool =>
+ (if sumbool_of_bool w__219 then
+ returnm (concat_vec ('b"1010001" : mword 7)
+ (concat_vec (rs2 : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_TYPE_D (rs1, rd, FCLASS_D) =>
+ (haveDExt tt) >>= fun w__222 : bool =>
+ (if sumbool_of_bool w__222 then
+ returnm (concat_vec ('b"1110001" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"001" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_TYPE_D (rs1, rd, FMV_X_D) =>
+ (is_RV64D tt) >>= fun w__225 : bool =>
+ (if sumbool_of_bool w__225 then
+ returnm (concat_vec ('b"1110001" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | F_UN_TYPE_D (rs1, rd, FMV_D_X) =>
+ (is_RV64D tt) >>= fun w__228 : bool =>
+ (if sumbool_of_bool w__228 then
+ returnm (concat_vec ('b"1111001" : mword 7)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec (rs1 : mword 5)
+ (concat_vec ('b"000" : mword 3)
+ (concat_vec (rd : mword 5) ('b"1010011" : mword 7))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 32)
+ | ILLEGAL s => returnm s
| _ => assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt
end)
: M (mword 32).
-Definition encdec_backwards (arg_ : mword 32)
-: M (ast) :=
-
- let v__7 := arg_ in
- let _mappingpatterns_23_ : mword 7 := subrange_vec_dec v__7 6 0 in
- (and_boolM (returnm ((encdec_uop_backwards_matches _mappingpatterns_23_) : bool))
- ((if ((encdec_uop_backwards_matches _mappingpatterns_23_)) then
- (encdec_uop_backwards _mappingpatterns_23_) >>= fun op => returnm ((true : bool) : bool)
- else returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))})))
+Definition encdec_backwards (arg_ : mword 32) : M (ast) :=
+ let v__9 := arg_ in
+ let _mappingpatterns_59_ : mword 7 := subrange_vec_dec v__9 6 0 in
+ (and_boolM (returnm ((encdec_uop_backwards_matches _mappingpatterns_59_) : bool))
+ ((if encdec_uop_backwards_matches _mappingpatterns_59_ then
+ (encdec_uop_backwards _mappingpatterns_59_) >>= fun op => returnm (true : bool)
+ else returnm false)
: M (bool))) >>= fun w__1 : bool =>
- (if sumbool_of_bool (w__1) then
- let imm : mword 20 := subrange_vec_dec v__7 31 12 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let imm : mword 20 := subrange_vec_dec v__7 31 12 in
- let _mappingpatterns_23_ : mword 7 := subrange_vec_dec v__7 6 0 in
- (encdec_uop_backwards _mappingpatterns_23_) >>= fun op =>
- returnm ((UTYPE
- ((imm, rd, op)))
- : ast)
- else if ((eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B1;B1;B0;B1;B1;B1;B1] : mword (6 - 0 + 1)))) then
- let imm_19 : bits 1 := subrange_vec_dec v__7 31 31 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let imm_8 : bits 1 := subrange_vec_dec v__7 20 20 in
- let imm_7_0 : bits 8 := subrange_vec_dec v__7 19 12 in
- let imm_19 : bits 1 := subrange_vec_dec v__7 31 31 in
- let imm_18_13 : bits 6 := subrange_vec_dec v__7 30 25 in
- let imm_12_9 : bits 4 := subrange_vec_dec v__7 24 21 in
- returnm ((RISCV_JAL
- ((concat_vec (imm_19 : bits 1)
- (concat_vec (imm_7_0 : bits 8)
- (concat_vec (imm_8 : bits 1)
- (concat_vec (imm_18_13 : bits 6)
- (concat_vec (imm_12_9 : bits 4) (vec_of_bits [B0] : mword 1))))), rd)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B1;B1;B0;B0;B1;B1;B1] : mword (6 - 0 + 1))))) then
- let imm : mword 12 := subrange_vec_dec v__7 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let imm : mword 12 := subrange_vec_dec v__7 31 20 in
- returnm ((RISCV_JALR
- ((imm, rs1, rd)))
- : ast)
+ (if sumbool_of_bool w__1 then
+ let imm : mword 20 := subrange_vec_dec v__9 31 12 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let imm : mword 20 := subrange_vec_dec v__9 31 12 in
+ let _mappingpatterns_59_ : mword 7 := subrange_vec_dec v__9 6 0 in
+ (encdec_uop_backwards _mappingpatterns_59_) >>= fun op => returnm (UTYPE (imm, rd, op))
+ else if eq_vec (subrange_vec_dec v__9 6 0) ('b"1101111" : mword (6 - 0 + 1)) then
+ let imm_19 : bits 1 := subrange_vec_dec v__9 31 31 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let imm_8 : bits 1 := subrange_vec_dec v__9 20 20 in
+ let imm_7_0 : bits 8 := subrange_vec_dec v__9 19 12 in
+ let imm_19 : bits 1 := subrange_vec_dec v__9 31 31 in
+ let imm_18_13 : bits 6 := subrange_vec_dec v__9 30 25 in
+ let imm_12_9 : bits 4 := subrange_vec_dec v__9 24 21 in
+ returnm (RISCV_JAL
+ (concat_vec (imm_19 : bits 1)
+ (concat_vec (imm_7_0 : bits 8)
+ (concat_vec (imm_8 : bits 1)
+ (concat_vec (imm_18_13 : bits 6)
+ (concat_vec (imm_12_9 : bits 4) ('b"0" : mword 1))))), rd))
+ else if andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"1100111" : mword (6 - 0 + 1))) then
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ returnm (RISCV_JALR (imm, rs1, rd))
else
(and_boolM
- (let _mappingpatterns_24_ : mword 3 := subrange_vec_dec v__7 14 12 in
- (and_boolM (returnm ((encdec_bop_backwards_matches _mappingpatterns_24_) : bool))
- ((if ((encdec_bop_backwards_matches _mappingpatterns_24_)) then
- (encdec_bop_backwards _mappingpatterns_24_) >>= fun op =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ (let _mappingpatterns_60_ : mword 3 := subrange_vec_dec v__9 14 12 in
+ (and_boolM (returnm ((encdec_bop_backwards_matches _mappingpatterns_60_) : bool))
+ ((if encdec_bop_backwards_matches _mappingpatterns_60_ then
+ (encdec_bop_backwards _mappingpatterns_60_) >>= fun op => returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
- (returnm ((eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B1;B1;B0;B0;B0;B1;B1] : mword (6 - 0 + 1)))
- : bool))) >>= fun w__4 : bool =>
- (if sumbool_of_bool (w__4) then
- let imm7_6 : bits 1 := subrange_vec_dec v__7 31 31 in
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let imm7_6 : bits 1 := subrange_vec_dec v__7 31 31 in
- let imm7_5_0 : bits 6 := subrange_vec_dec v__7 30 25 in
- let imm5_4_1 : bits 4 := subrange_vec_dec v__7 11 8 in
- let imm5_0 : bits 1 := subrange_vec_dec v__7 7 7 in
- let _mappingpatterns_24_ : mword 3 := subrange_vec_dec v__7 14 12 in
- (encdec_bop_backwards _mappingpatterns_24_) >>= fun op =>
- returnm ((BTYPE
- ((concat_vec (imm7_6 : bits 1)
- (concat_vec (imm5_0 : bits 1)
- (concat_vec (imm7_5_0 : bits 6)
- (concat_vec (imm5_4_1 : bits 4) (vec_of_bits [B0] : mword 1)))), rs2, rs1, op)))
- : ast)
+ (returnm ((eq_vec (subrange_vec_dec v__9 6 0) ('b"1100011" : mword (6 - 0 + 1))) : bool))) >>= fun w__4 : bool =>
+ (if sumbool_of_bool w__4 then
+ let imm7_6 : bits 1 := subrange_vec_dec v__9 31 31 in
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let imm7_6 : bits 1 := subrange_vec_dec v__9 31 31 in
+ let imm7_5_0 : bits 6 := subrange_vec_dec v__9 30 25 in
+ let imm5_4_1 : bits 4 := subrange_vec_dec v__9 11 8 in
+ let imm5_0 : bits 1 := subrange_vec_dec v__9 7 7 in
+ let _mappingpatterns_60_ : mword 3 := subrange_vec_dec v__9 14 12 in
+ (encdec_bop_backwards _mappingpatterns_60_) >>= fun op =>
+ returnm (BTYPE
+ (concat_vec (imm7_6 : bits 1)
+ (concat_vec (imm5_0 : bits 1)
+ (concat_vec (imm7_5_0 : bits 6)
+ (concat_vec (imm5_4_1 : bits 4) ('b"0" : mword 1)))), rs2, rs1, op))
else
(and_boolM
- (let _mappingpatterns_25_ : mword 3 := subrange_vec_dec v__7 14 12 in
- (and_boolM (returnm ((encdec_iop_backwards_matches _mappingpatterns_25_) : bool))
- ((if ((encdec_iop_backwards_matches _mappingpatterns_25_)) then
- (encdec_iop_backwards _mappingpatterns_25_) >>= fun op =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ (let _mappingpatterns_61_ : mword 3 := subrange_vec_dec v__9 14 12 in
+ (and_boolM (returnm ((encdec_iop_backwards_matches _mappingpatterns_61_) : bool))
+ ((if encdec_iop_backwards_matches _mappingpatterns_61_ then
+ (encdec_iop_backwards _mappingpatterns_61_) >>= fun op => returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
- (returnm ((eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword (6 - 0 + 1)))
+ (returnm ((eq_vec (subrange_vec_dec v__9 6 0) ('b"0010011" : mword (6 - 0 + 1)))
: bool))) >>= fun w__7 : bool =>
- (if sumbool_of_bool (w__7) then
- let imm : mword 12 := subrange_vec_dec v__7 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let imm : mword 12 := subrange_vec_dec v__7 31 20 in
- let _mappingpatterns_25_ : mword 3 := subrange_vec_dec v__7 14 12 in
- (encdec_iop_backwards _mappingpatterns_25_) >>= fun op =>
- returnm ((ITYPE
- ((imm, rs1, rd, op)))
- : ast)
+ (if sumbool_of_bool w__7 then
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ let _mappingpatterns_61_ : mword 3 := subrange_vec_dec v__9 14 12 in
+ (encdec_iop_backwards _mappingpatterns_61_) >>= fun op =>
+ returnm (ITYPE (imm, rs1, rd, op))
+ else if sumbool_of_bool
+ (andb
+ (let shamt : mword 6 := subrange_vec_dec v__9 25 20 in
+ orb (Z.eqb 64 64) (eq_bit (access_vec_dec shamt 5) B0))
+ (andb (eq_vec (subrange_vec_dec v__9 31 26) ('b"000000" : mword (31 - 26 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0010011" : mword (6 - 0 + 1)))))) then
+ let shamt : mword 6 := subrange_vec_dec v__9 25 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTIOP (shamt, rs1, rd, RISCV_SLLI))
+ else if sumbool_of_bool
+ (andb
+ (let shamt : mword 6 := subrange_vec_dec v__9 25 20 in
+ orb (Z.eqb 64 64) (eq_bit (access_vec_dec shamt 5) B0))
+ (andb (eq_vec (subrange_vec_dec v__9 31 26) ('b"000000" : mword (31 - 26 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0010011" : mword (6 - 0 + 1)))))) then
+ let shamt : mword 6 := subrange_vec_dec v__9 25 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTIOP (shamt, rs1, rd, RISCV_SRLI))
+ else if sumbool_of_bool
+ (andb
+ (let shamt : mword 6 := subrange_vec_dec v__9 25 20 in
+ orb (Z.eqb 64 64) (eq_bit (access_vec_dec shamt 5) B0))
+ (andb (eq_vec (subrange_vec_dec v__9 31 26) ('b"010000" : mword (31 - 26 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0010011" : mword (6 - 0 + 1)))))) then
+ let shamt : mword 6 := subrange_vec_dec v__9 25 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTIOP (shamt, rs1, rd, RISCV_SRAI))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_ADD))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"010" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_SLT))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"011" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_SLTU))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"111" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_AND))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"110" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_OR))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"100" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_XOR))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_SLL))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_SRL))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_SUB))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0110011" : mword (6 - 0 + 1)))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPE (rs2, rs1, rd, RISCV_SRA))
else
- (and_boolMP
- ((let shamt : mword 6 := subrange_vec_dec v__7 25 20 in
- (or_boolMP
- ((returnm (build_ex
- (projT1
- (build_ex
- (Z.eqb 64 64)
- : {_bool : bool & ArithFact (iff (_bool = true) (64 = 64))})))) : M ({_bool : bool & ArithFact (iff (_bool =
- true) (64 = 64))}))
- (build_trivial_ex
- ((bit_to_bool (access_vec_dec shamt 5)) >>= fun w__8 : bool =>
- returnm ((Bool.eqb w__8 false)
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 =
- 64 \/
- simp_0 = true))}))
- : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 = 64 \/
- simp_0 = true))})) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (64 = 64 \/ simp_0 = true))}))
- (build_trivial_ex
- (returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 26)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (31 - 26 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword (6 - 0 + 1)))))
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool = true) (simp_0 =
- true /\
- (64 = 64 \/ simp_1 = true)))})) >>= fun '(existT _ w__10 _ : {_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))}) =>
- (if sumbool_of_bool (w__10) then
- let shamt : mword 6 := subrange_vec_dec v__7 25 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTIOP
- ((shamt, rs1, rd, RISCV_SLLI)))
- : ast)
+ (and_boolM
+ (let _mappingpatterns_63_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_62_ : mword 1 := subrange_vec_dec v__9 14 14 in
+ (and_boolM (returnm ((size_bits_backwards_matches _mappingpatterns_63_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_63_ then
+ (size_bits_backwards _mappingpatterns_63_) >>= fun size =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_62_) : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_62_ then
+ (bool_bits_backwards _mappingpatterns_62_) >>= fun is_unsigned =>
+ returnm ((orb (Z.ltb (projT1 (word_width_bytes size)) 8)
+ (andb (negb is_unsigned)
+ ((Z.leb (projT1 (word_width_bytes size)) 8)
+ : bool)))
+ : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((eq_vec (subrange_vec_dec v__9 6 0) ('b"0000011" : mword (6 - 0 + 1)))
+ : bool))) >>= fun w__12 : bool =>
+ (if sumbool_of_bool w__12 then
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ let _mappingpatterns_63_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_62_ : mword 1 := subrange_vec_dec v__9 14 14 in
+ (size_bits_backwards _mappingpatterns_63_) >>= fun size =>
+ (bool_bits_backwards _mappingpatterns_62_) >>= fun is_unsigned =>
+ returnm (LOAD (imm, rs1, rd, is_unsigned, size, false, false))
else
- (and_boolMP
- ((let shamt : mword 6 := subrange_vec_dec v__7 25 20 in
- (or_boolMP
- ((returnm (build_ex
- (projT1
- (build_ex
- (Z.eqb 64 64)
- : {_bool : bool & ArithFact (iff (_bool = true) (64 = 64))})))) : M ({_bool : bool & ArithFact (iff (_bool =
- true) (64 = 64))}))
- (build_trivial_ex
- ((bit_to_bool (access_vec_dec shamt 5)) >>= fun w__11 : bool =>
- returnm ((Bool.eqb w__11 false)
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 =
- 64 \/
- simp_0 = true))}))
- : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 = 64 \/
- simp_0 = true))})) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (64 = 64 \/ simp_0 = true))}))
- (build_trivial_ex
+ (and_boolM
+ (let _mappingpatterns_64_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ (and_boolM (returnm ((size_bits_backwards_matches _mappingpatterns_64_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_64_ then
+ (size_bits_backwards _mappingpatterns_64_) >>= fun size =>
+ returnm ((Z.leb (projT1 (word_width_bytes size)) 8) : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 26)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (31 - 26 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword (6 - 0 + 1)))))
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))})) >>= fun '(existT _ w__13 _ : {_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))}) =>
- (if sumbool_of_bool (w__13) then
- let shamt : mword 6 := subrange_vec_dec v__7 25 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTIOP
- ((shamt, rs1, rd, RISCV_SRLI)))
- : ast)
+ (eq_vec (subrange_vec_dec v__9 14 14) ('b"0" : mword (14 - 14 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0100011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__15 : bool =>
+ (if sumbool_of_bool w__15 then
+ let imm7 : bits 7 := subrange_vec_dec v__9 31 25 in
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let imm7 : bits 7 := subrange_vec_dec v__9 31 25 in
+ let imm5 : bits 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_64_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ (size_bits_backwards _mappingpatterns_64_) >>= fun size =>
+ returnm (STORE
+ (concat_vec (imm7 : bits 7) (imm5 : bits 5), rs2, rs1, size, false, false))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__9 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0011011" : mword (6 - 0 + 1)))))
+ then
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ returnm (ADDIW (imm, rs1, rd))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ let shamt : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTW (shamt, rs1, rd, RISCV_SLLI))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ let shamt : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTW (shamt, rs1, rd, RISCV_SRLI))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ let shamt : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTW (shamt, rs1, rd, RISCV_SRAI))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPEW (rs2, rs1, rd, RISCV_ADDW))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPEW (rs2, rs1, rd, RISCV_SUBW))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPEW (rs2, rs1, rd, RISCV_SLLW))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPEW (rs2, rs1, rd, RISCV_SRLW))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (RTYPEW (rs2, rs1, rd, RISCV_SRAW))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ let shamt : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTIWOP (shamt, rs1, rd, RISCV_SLLIW))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ let shamt : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTIWOP (shamt, rs1, rd, RISCV_SRLIW))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ let shamt : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (SHIFTIWOP (shamt, rs1, rd, RISCV_SRAIW))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 28) (Ox"0" : mword (31 - 28 + 1)))
+ (eq_vec (subrange_vec_dec v__9 19 0) (Ox"0000F" : mword (19 - 0 + 1)))
+ then
+ let succ : mword 4 := subrange_vec_dec v__9 23 20 in
+ let pred : mword 4 := subrange_vec_dec v__9 27 24 in
+ returnm (FENCE (pred, succ))
+ else if andb (eq_vec (subrange_vec_dec v__9 31 28) (Ox"8" : mword (31 - 28 + 1)))
+ (eq_vec (subrange_vec_dec v__9 19 0) (Ox"0000F" : mword (19 - 0 + 1)))
+ then
+ let succ : mword 4 := subrange_vec_dec v__9 23 20 in
+ let pred : mword 4 := subrange_vec_dec v__9 27 24 in
+ returnm (FENCE_TSO (pred, succ))
+ else if eq_vec v__9 (Ox"0000100F" : mword 32) then returnm (FENCEI tt)
+ else if eq_vec v__9 (Ox"00000073" : mword 32) then returnm (ECALL tt)
+ else if eq_vec v__9 (Ox"30200073" : mword 32) then returnm (MRET tt)
+ else if eq_vec v__9 (Ox"10200073" : mword 32) then returnm (SRET tt)
+ else if eq_vec v__9 (Ox"00100073" : mword 32) then returnm (EBREAK tt)
+ else if eq_vec v__9 (Ox"10500073" : mword 32) then returnm (WFI tt)
+ else if andb
+ (eq_vec (subrange_vec_dec v__9 31 25) ('b"0001001" : mword (31 - 25 + 1)))
+ (eq_vec (subrange_vec_dec v__9 14 0)
+ ('b"000000001110011"
+ : mword (14 - 0 + 1))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ returnm (SFENCE_VMA (rs1, rs2))
else
- (and_boolMP
- ((let shamt : mword 6 := subrange_vec_dec v__7 25 20 in
- (or_boolMP
- ((returnm (build_ex
- (projT1
- (build_ex
- (Z.eqb 64 64)
- : {_bool : bool & ArithFact (iff (_bool = true) (64 = 64))})))) : M ({_bool : bool & ArithFact (iff (_bool =
- true) (64 = 64))}))
- (build_trivial_ex
- ((bit_to_bool (access_vec_dec shamt 5)) >>= fun w__14 : bool =>
- returnm ((Bool.eqb w__14 false)
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (64 = 64 \/ simp_0 = true))}))
- : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 = 64 \/
- simp_0 = true))})) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (64 = 64 \/ simp_0 = true))}))
- (build_trivial_ex
+ (and_boolM
+ (let _mappingpatterns_67_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_66_ : mword 1 := subrange_vec_dec v__9 25 25 in
+ let _mappingpatterns_65_ : mword 1 := subrange_vec_dec v__9 26 26 in
+ (and_boolM
+ (returnm ((size_bits_backwards_matches _mappingpatterns_67_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_67_ then
+ (size_bits_backwards _mappingpatterns_67_) >>= fun size =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_66_) : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_66_ then
+ (bool_bits_backwards _mappingpatterns_66_) >>= fun rl =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_65_)
+ : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_65_ then
+ (bool_bits_backwards _mappingpatterns_65_) >>= fun aq =>
+ returnm ((Z.leb (projT1 (word_width_bytes size)) 8)
+ : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 26)
- (vec_of_bits [B0;B1;B0;B0;B0;B0] : mword (31 - 26 + 1)))
+ (eq_vec (subrange_vec_dec v__9 31 27)
+ ('b"00010"
+ : mword (31 - 27 + 1)))
(andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword (6 - 0 + 1)))))
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))})) >>= fun '(existT _ w__16 _ : {_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))}) =>
- (if sumbool_of_bool (w__16) then
- let shamt : mword 6 := subrange_vec_dec v__7 25 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTIOP
- ((shamt, rs1, rd, RISCV_SRAI)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_ADD)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B1;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_SLT)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B1;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_SLTU)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B1;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_AND)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B1;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_OR)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_XOR)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_SLL)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_SRL)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_SUB)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPE
- ((rs2, rs1, rd, RISCV_SRA)))
- : ast)
+ (eq_vec (subrange_vec_dec v__9 24 20)
+ ('b"00000"
+ : mword (24 - 20 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 14)
+ ('b"0"
+ : mword (14 - 14 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0101111"
+ : mword (6 - 0 + 1))))))
+ : bool))) >>= fun w__22 : bool =>
+ (if sumbool_of_bool w__22 then
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_67_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_66_ : mword 1 := subrange_vec_dec v__9 25 25 in
+ let _mappingpatterns_65_ : mword 1 := subrange_vec_dec v__9 26 26 in
+ (size_bits_backwards _mappingpatterns_67_) >>= fun size =>
+ (bool_bits_backwards _mappingpatterns_66_) >>= fun rl =>
+ (bool_bits_backwards _mappingpatterns_65_) >>= fun aq =>
+ returnm (LOADRES (aq, rl, rs1, size, rd))
else
(and_boolM
- (let _mappingpatterns_27_ : mword 2 := subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_26_ : mword 1 := subrange_vec_dec v__7 14 14 in
+ (let _mappingpatterns_70_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_69_ : mword 1 := subrange_vec_dec v__9 25 25 in
+ let _mappingpatterns_68_ : mword 1 := subrange_vec_dec v__9 26 26 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_27_)
- : bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_27_)) then
- (size_bits_backwards _mappingpatterns_27_) >>= fun size =>
+ (returnm ((size_bits_backwards_matches _mappingpatterns_70_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_70_ then
+ (size_bits_backwards _mappingpatterns_70_) >>= fun size =>
(and_boolM
- (returnm ((bool_bits_backwards_matches _mappingpatterns_26_)
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_69_)
: bool))
- ((if ((bool_bits_backwards_matches _mappingpatterns_26_)) then
- (bool_bits_backwards _mappingpatterns_26_) >>= fun is_unsigned =>
- returnm (((orb (Z.ltb (projT1 (word_width_bytes size)) 8)
- (andb (negb is_unsigned)
- ((Z.leb (projT1 (word_width_bytes size)) 8)
- : bool)))
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__18 : bool =>
- returnm ((w__18
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if bool_bits_backwards_matches _mappingpatterns_69_ then
+ (bool_bits_backwards _mappingpatterns_69_) >>= fun rl =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_68_)
+ : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_68_ then
+ (bool_bits_backwards _mappingpatterns_68_) >>= fun aq =>
+ returnm ((Z.leb (projT1 (word_width_bytes size)) 8)
+ : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
: M (bool)))
: M (bool))
- (returnm ((eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B1;B1] : mword (6 - 0 + 1)))
- : bool))) >>= fun w__21 : bool =>
- (if sumbool_of_bool (w__21) then
- let imm : mword 12 := subrange_vec_dec v__7 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let imm : mword 12 := subrange_vec_dec v__7 31 20 in
- let _mappingpatterns_27_ : mword 2 := subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_26_ : mword 1 := subrange_vec_dec v__7 14 14 in
- (size_bits_backwards _mappingpatterns_27_) >>= fun size =>
- (bool_bits_backwards _mappingpatterns_26_) >>= fun is_unsigned =>
- returnm ((LOAD
- ((imm, rs1, rd, is_unsigned, size, false, false)))
- : ast)
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__9 31 27)
+ ('b"00011"
+ : mword (31 - 27 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 14)
+ ('b"0"
+ : mword (14 - 14 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0101111"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__29 : bool =>
+ (if sumbool_of_bool w__29 then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_70_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_69_ : mword 1 := subrange_vec_dec v__9 25 25 in
+ let _mappingpatterns_68_ : mword 1 := subrange_vec_dec v__9 26 26 in
+ (size_bits_backwards _mappingpatterns_70_) >>= fun size =>
+ (bool_bits_backwards _mappingpatterns_69_) >>= fun rl =>
+ (bool_bits_backwards _mappingpatterns_68_) >>= fun aq =>
+ returnm (STORECON (aq, rl, rs2, rs1, size, rd))
else
(and_boolM
- (let _mappingpatterns_28_ : mword 2 := subrange_vec_dec v__7 13 12 in
+ (let _mappingpatterns_71_ : mword 5 := subrange_vec_dec v__9 31 27 in
+ let _mappingpatterns_74_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_73_ : mword 1 := subrange_vec_dec v__9 25 25 in
+ let _mappingpatterns_72_ : mword 1 := subrange_vec_dec v__9 26 26 in
+ let _mappingpatterns_71_ : mword 5 := subrange_vec_dec v__9 31 27 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_28_)
- : bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_28_)) then
- (size_bits_backwards _mappingpatterns_28_) >>= fun size =>
- returnm (((Z.leb (projT1 (word_width_bytes size)) 8)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ (returnm ((size_bits_backwards_matches _mappingpatterns_74_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_74_ then
+ (size_bits_backwards _mappingpatterns_74_) >>= fun size =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_73_)
+ : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_73_ then
+ (bool_bits_backwards _mappingpatterns_73_) >>= fun rl =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches
+ _mappingpatterns_72_)
+ : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_72_
+ then
+ (bool_bits_backwards _mappingpatterns_72_) >>= fun aq =>
+ (and_boolM
+ (returnm ((encdec_amoop_backwards_matches
+ _mappingpatterns_71_)
+ : bool))
+ ((if encdec_amoop_backwards_matches
+ _mappingpatterns_71_ then
+ (encdec_amoop_backwards _mappingpatterns_71_) >>= fun op =>
+ returnm ((Z.leb
+ (projT1
+ (word_width_bytes size)) 8)
+ : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 14 14)
- (vec_of_bits [B0] : mword (14 - 14 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B0;B0;B0;B1;B1] : mword (6 - 0 + 1))))
- : bool))) >>= fun w__24 : bool =>
- (if sumbool_of_bool (w__24) then
- let imm7 : bits 7 := subrange_vec_dec v__7 31 25 in
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let imm7 : bits 7 := subrange_vec_dec v__7 31 25 in
- let imm5 : bits 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_28_ : mword 2 := subrange_vec_dec v__7 13 12 in
- (size_bits_backwards _mappingpatterns_28_) >>= fun size =>
- returnm ((STORE
- ((concat_vec (imm7 : bits 7) (imm5 : bits 5), rs2, rs1, size, false, false)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B0]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1)))))) then
- let imm : mword 12 := subrange_vec_dec v__7 31 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let imm : mword 12 := subrange_vec_dec v__7 31 20 in
- returnm ((ADDIW
- ((imm, rs1, rd)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let shamt : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTW
- ((shamt, rs1, rd, RISCV_SLLI)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let shamt : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTW
- ((shamt, rs1, rd, RISCV_SRLI)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let shamt : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTW
- ((shamt, rs1, rd, RISCV_SRAI)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B0]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPEW
- ((rs2, rs1, rd, RISCV_ADDW)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B0]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPEW
- ((rs2, rs1, rd, RISCV_SUBW)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPEW
- ((rs2, rs1, rd, RISCV_SLLW)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPEW
- ((rs2, rs1, rd, RISCV_SRLW)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((RTYPEW
- ((rs2, rs1, rd, RISCV_SRAW)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B0;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let shamt : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTIWOP
- ((shamt, rs1, rd, RISCV_SLLIW)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let shamt : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTIWOP
- ((shamt, rs1, rd, RISCV_SRLIW)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- let shamt : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((SHIFTIWOP
- ((shamt, rs1, rd, RISCV_SRAIW)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 28)
- (vec_of_bits [B0;B0;B0;B0] : mword (31 - 28 + 1)))
- (eq_vec (subrange_vec_dec v__7 19 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B1]
- : mword (19 - 0 + 1))))) then
- let succ : mword 4 := subrange_vec_dec v__7 23 20 in
- let pred : mword 4 := subrange_vec_dec v__7 27 24 in
- returnm ((FENCE
- ((pred, succ)))
- : ast)
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 28)
- (vec_of_bits [B1;B0;B0;B0] : mword (31 - 28 + 1)))
- (eq_vec (subrange_vec_dec v__7 19 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B1]
- : mword (19 - 0 + 1))))) then
- let succ : mword 4 := subrange_vec_dec v__7 23 20 in
- let pred : mword 4 := subrange_vec_dec v__7 27 24 in
- returnm ((FENCE_TSO
- ((pred, succ)))
- : ast)
- else if ((eq_vec v__7
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;
- B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B1]
- : mword 32))) then
- returnm ((FENCEI
- (tt))
- : ast )
- else if ((eq_vec v__7
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm ((ECALL
- (tt))
- : ast )
- else if ((eq_vec v__7
- (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm ((MRET
- (tt))
- : ast )
- else if ((eq_vec v__7
- (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm ((SRET
- (tt))
- : ast )
- else if ((eq_vec v__7
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm ((EBREAK
- (tt))
- : ast )
- else if ((eq_vec v__7
- (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm ((WFI
- (tt))
- : ast )
- else if ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B1;B0;B0;B1] : mword (31 - 25 + 1)))
- (eq_vec (subrange_vec_dec v__7 14 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword (14 - 0 + 1))))) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- returnm ((SFENCE_VMA
- ((rs1, rs2)))
- : ast)
+ (eq_vec (subrange_vec_dec v__9 14 14)
+ ('b"0"
+ : mword (14 - 14 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0101111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__38 : bool =>
+ (if sumbool_of_bool w__38 then
+ let _mappingpatterns_71_ : mword 5 := subrange_vec_dec v__9 31 27 in
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_74_ : mword 2 := subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_73_ : mword 1 := subrange_vec_dec v__9 25 25 in
+ let _mappingpatterns_72_ : mword 1 := subrange_vec_dec v__9 26 26 in
+ let _mappingpatterns_71_ : mword 5 := subrange_vec_dec v__9 31 27 in
+ (size_bits_backwards _mappingpatterns_74_) >>= fun size =>
+ (bool_bits_backwards _mappingpatterns_73_) >>= fun rl =>
+ (bool_bits_backwards _mappingpatterns_72_) >>= fun aq =>
+ (encdec_amoop_backwards _mappingpatterns_71_) >>= fun op =>
+ returnm (AMO (op, aq, rl, rs2, rs1, size, rd))
else
(and_boolM
- (let _mappingpatterns_31_ : mword 2 := subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_30_ : mword 1 := subrange_vec_dec v__7 25 25 in
- let _mappingpatterns_29_ : mword 1 := subrange_vec_dec v__7 26 26 in
+ (let _mappingpatterns_75_ : mword 3 := subrange_vec_dec v__9 14 12 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_31_)
+ (returnm ((encdec_mul_op_backwards_matches _mappingpatterns_75_)
: bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_31_)) then
- (size_bits_backwards _mappingpatterns_31_) >>= fun size =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches _mappingpatterns_30_)
- : bool))
- ((if ((bool_bits_backwards_matches _mappingpatterns_30_))
- then
- (bool_bits_backwards _mappingpatterns_30_) >>= fun rl =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_29_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_29_)) then
- (bool_bits_backwards _mappingpatterns_29_) >>= fun aq =>
- returnm (((Z.leb (projT1 (word_width_bytes size))
- 8)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__26 : bool =>
- returnm ((w__26
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__28 : bool =>
- returnm ((w__28
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if encdec_mul_op_backwards_matches _mappingpatterns_75_ then
+ (encdec_mul_op_backwards _mappingpatterns_75_) >>= fun '(high, signed1, signed2) =>
+ returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 27)
- (vec_of_bits [B0;B0;B0;B1;B0] : mword (31 - 27 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 24 20)
- (vec_of_bits [B0;B0;B0;B0;B0]
- : mword (24 - 20 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 14)
- (vec_of_bits [B0] : mword (14 - 14 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1]
- : mword (6 - 0 + 1))))))
- : bool))) >>= fun w__31 : bool =>
- (if sumbool_of_bool (w__31) then
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_31_ : mword 2 := subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_30_ : mword 1 := subrange_vec_dec v__7 25 25 in
- let _mappingpatterns_29_ : mword 1 := subrange_vec_dec v__7 26 26 in
- (size_bits_backwards _mappingpatterns_31_) >>= fun size =>
- (bool_bits_backwards _mappingpatterns_30_) >>= fun rl =>
- (bool_bits_backwards _mappingpatterns_29_) >>= fun aq =>
- returnm ((LOADRES
- ((aq, rl, rs1, size, rd)))
- : ast)
+ (eq_vec (subrange_vec_dec v__9 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0110011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__41 : bool =>
+ (if sumbool_of_bool w__41 then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_75_ : mword 3 := subrange_vec_dec v__9 14 12 in
+ (encdec_mul_op_backwards _mappingpatterns_75_) >>= fun '(high, signed1, signed2) =>
+ returnm (MUL (rs2, rs1, rd, high, signed1, signed2))
else
(and_boolM
- (let _mappingpatterns_34_ : mword 2 := subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_33_ : mword 1 := subrange_vec_dec v__7 25 25 in
- let _mappingpatterns_32_ : mword 1 := subrange_vec_dec v__7 26 26 in
+ (let _mappingpatterns_76_ : mword 1 := subrange_vec_dec v__9 12 12 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_34_)
+ (returnm ((bool_not_bits_backwards_matches _mappingpatterns_76_)
: bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_34_)) then
- (size_bits_backwards _mappingpatterns_34_) >>= fun size =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_33_)
- : bool))
- ((if ((bool_bits_backwards_matches _mappingpatterns_33_))
- then
- (bool_bits_backwards _mappingpatterns_33_) >>= fun rl =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_32_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_32_)) then
- (bool_bits_backwards _mappingpatterns_32_) >>= fun aq =>
- returnm (((Z.leb
- (projT1
- (word_width_bytes size)) 8)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool))) >>= fun w__33 : bool =>
- returnm ((w__33
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__35 : bool =>
- returnm ((w__35
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if bool_not_bits_backwards_matches _mappingpatterns_76_ then
+ (bool_not_bits_backwards _mappingpatterns_76_) >>= fun s =>
+ returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 27)
- (vec_of_bits [B0;B0;B0;B1;B1]
- : mword (31 - 27 + 1)))
+ (eq_vec (subrange_vec_dec v__9 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
(andb
- (eq_vec (subrange_vec_dec v__7 14 14)
- (vec_of_bits [B0] : mword (14 - 14 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1]
- : mword (6 - 0 + 1)))))
- : bool))) >>= fun w__38 : bool =>
- (if sumbool_of_bool (w__38) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_34_ : mword 2 := subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_33_ : mword 1 := subrange_vec_dec v__7 25 25 in
- let _mappingpatterns_32_ : mword 1 := subrange_vec_dec v__7 26 26 in
- (size_bits_backwards _mappingpatterns_34_) >>= fun size =>
- (bool_bits_backwards _mappingpatterns_33_) >>= fun rl =>
- (bool_bits_backwards _mappingpatterns_32_) >>= fun aq =>
- returnm ((STORECON
- ((aq, rl, rs2, rs1, size, rd)))
- : ast)
+ (eq_vec (subrange_vec_dec v__9 14 13)
+ ('b"10"
+ : mword (14 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0110011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__44 : bool =>
+ (if sumbool_of_bool w__44 then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_76_ : mword 1 := subrange_vec_dec v__9 12 12 in
+ (bool_not_bits_backwards _mappingpatterns_76_) >>= fun s =>
+ returnm (DIV (rs2, rs1, rd, s))
else
(and_boolM
- (let _mappingpatterns_35_ : mword 5 :=
- subrange_vec_dec v__7 31 27 in
- let _mappingpatterns_38_ : mword 2 :=
- subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_37_ : mword 1 :=
- subrange_vec_dec v__7 25 25 in
- let _mappingpatterns_36_ : mword 1 :=
- subrange_vec_dec v__7 26 26 in
- let _mappingpatterns_35_ : mword 5 :=
- subrange_vec_dec v__7 31 27 in
+ (let _mappingpatterns_77_ : mword 1 :=
+ subrange_vec_dec v__9 12 12 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_38_)
+ (returnm ((bool_not_bits_backwards_matches
+ _mappingpatterns_77_)
: bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_38_))
+ ((if bool_not_bits_backwards_matches _mappingpatterns_77_
then
- (size_bits_backwards _mappingpatterns_38_) >>= fun size =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_37_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_37_)) then
- (bool_bits_backwards _mappingpatterns_37_) >>= fun rl =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_36_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_36_)) then
- (bool_bits_backwards _mappingpatterns_36_) >>= fun aq =>
- (and_boolM
- (returnm ((encdec_amoop_backwards_matches
- _mappingpatterns_35_)
- : bool))
- ((if ((encdec_amoop_backwards_matches
- _mappingpatterns_35_)) then
- (encdec_amoop_backwards
- _mappingpatterns_35_) >>= fun op =>
- returnm (((Z.leb
- (projT1
- (word_width_bytes size))
- 8)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool))) >>= fun w__40 : bool =>
- returnm ((w__40
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool))) >>= fun w__42 : bool =>
- returnm ((w__42
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__44 : bool =>
- returnm ((w__44
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ (bool_not_bits_backwards _mappingpatterns_77_) >>= fun s =>
+ returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 14 14)
- (vec_of_bits [B0] : mword (14 - 14 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1]
- : mword (6 - 0 + 1))))
+ (eq_vec (subrange_vec_dec v__9 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 13)
+ ('b"11"
+ : mword (14 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0110011"
+ : mword (6 - 0 + 1)))))
: bool))) >>= fun w__47 : bool =>
- (if sumbool_of_bool (w__47) then
- let _mappingpatterns_35_ : mword 5 :=
- subrange_vec_dec v__7 31 27 in
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_38_ : mword 2 :=
- subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_37_ : mword 1 :=
- subrange_vec_dec v__7 25 25 in
- let _mappingpatterns_36_ : mword 1 :=
- subrange_vec_dec v__7 26 26 in
- let _mappingpatterns_35_ : mword 5 :=
- subrange_vec_dec v__7 31 27 in
- (size_bits_backwards _mappingpatterns_38_) >>= fun size =>
- (bool_bits_backwards _mappingpatterns_37_) >>= fun rl =>
- (bool_bits_backwards _mappingpatterns_36_) >>= fun aq =>
- (encdec_amoop_backwards _mappingpatterns_35_) >>= fun op =>
- returnm ((AMO
- ((op, aq, rl, rs2, rs1, size, rd)))
- : ast)
+ (if sumbool_of_bool w__47 then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_77_ : mword 1 :=
+ subrange_vec_dec v__9 12 12 in
+ (bool_not_bits_backwards _mappingpatterns_77_) >>= fun s =>
+ returnm (REM (rs2, rs1, rd, s))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__9 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0111011"
+ : mword (6 - 0 + 1)))))) then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ returnm (MULW (rs2, rs1, rd))
else
(and_boolM
- (let _mappingpatterns_39_ : mword 3 :=
- subrange_vec_dec v__7 14 12 in
+ (let _mappingpatterns_78_ : mword 1 :=
+ subrange_vec_dec v__9 12 12 in
(and_boolM
- (returnm ((encdec_mul_op_backwards_matches
- _mappingpatterns_39_)
+ (returnm ((bool_not_bits_backwards_matches
+ _mappingpatterns_78_)
: bool))
- ((if ((encdec_mul_op_backwards_matches
- _mappingpatterns_39_)) then
- (encdec_mul_op_backwards _mappingpatterns_39_) >>= fun '(high, signed1, signed2) =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if bool_not_bits_backwards_matches _mappingpatterns_78_
+ then
+ (bool_not_bits_backwards _mappingpatterns_78_) >>= fun s =>
+ returnm ((Z.eqb 64 64) : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1]
- : mword (6 - 0 + 1))))
+ (eq_vec (subrange_vec_dec v__9 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__9 14 13)
+ ('b"10"
+ : mword (14 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0111011"
+ : mword (6 - 0 + 1)))))
: bool))) >>= fun w__50 : bool =>
- (if sumbool_of_bool (w__50) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_39_ : mword 3 :=
- subrange_vec_dec v__7 14 12 in
- (encdec_mul_op_backwards _mappingpatterns_39_) >>= fun '(high, signed1, signed2) =>
- returnm ((MUL
- ((rs2, rs1, rd, high, signed1, signed2)))
- : ast)
+ (if sumbool_of_bool w__50 then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_78_ : mword 1 :=
+ subrange_vec_dec v__9 12 12 in
+ (bool_not_bits_backwards _mappingpatterns_78_) >>= fun s =>
+ returnm (DIVW (rs2, rs1, rd, s))
else
(and_boolM
- (let _mappingpatterns_40_ : mword 1 :=
- subrange_vec_dec v__7 12 12 in
+ (let _mappingpatterns_79_ : mword 1 :=
+ subrange_vec_dec v__9 12 12 in
(and_boolM
(returnm ((bool_not_bits_backwards_matches
- _mappingpatterns_40_)
+ _mappingpatterns_79_)
: bool))
- ((if ((bool_not_bits_backwards_matches
- _mappingpatterns_40_)) then
- (bool_not_bits_backwards _mappingpatterns_40_) >>= fun s =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if bool_not_bits_backwards_matches
+ _mappingpatterns_79_ then
+ (bool_not_bits_backwards _mappingpatterns_79_) >>= fun s =>
+ returnm ((Z.eqb 64 64) : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
+ (eq_vec (subrange_vec_dec v__9 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
(andb
- (eq_vec (subrange_vec_dec v__7 14 13)
- (vec_of_bits [B1;B0]
- : mword (14 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1]
- : mword (6 - 0 + 1)))))
+ (eq_vec (subrange_vec_dec v__9 14 13)
+ ('b"11"
+ : mword (14 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0111011"
+ : mword (6 - 0 + 1)))))
: bool))) >>= fun w__53 : bool =>
- (if sumbool_of_bool (w__53) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_40_ : mword 1 :=
- subrange_vec_dec v__7 12 12 in
- (bool_not_bits_backwards _mappingpatterns_40_) >>= fun s =>
- returnm ((DIV
- ((rs2, rs1, rd, s)))
- : ast)
+ (if sumbool_of_bool w__53 then
+ let rs2 : mword 5 := subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_79_ : mword 1 :=
+ subrange_vec_dec v__9 12 12 in
+ (bool_not_bits_backwards _mappingpatterns_79_) >>= fun s =>
+ returnm (REMW (rs2, rs1, rd, s))
else
(and_boolM
- (let _mappingpatterns_41_ : mword 1 :=
- subrange_vec_dec v__7 12 12 in
+ (let _mappingpatterns_81_ : mword 2 :=
+ subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_80_ : mword 1 :=
+ subrange_vec_dec v__9 14 14 in
(and_boolM
- (returnm ((bool_not_bits_backwards_matches
- _mappingpatterns_41_)
+ (returnm ((encdec_csrop_backwards_matches
+ _mappingpatterns_81_)
: bool))
- ((if ((bool_not_bits_backwards_matches
- _mappingpatterns_41_)) then
- (bool_not_bits_backwards _mappingpatterns_41_) >>= fun s =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if encdec_csrop_backwards_matches
+ _mappingpatterns_81_ then
+ (encdec_csrop_backwards _mappingpatterns_81_) >>= fun op =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches
+ _mappingpatterns_80_)
+ : bool))
+ ((if bool_bits_backwards_matches
+ _mappingpatterns_80_ then
+ (bool_bits_backwards _mappingpatterns_80_) >>= fun is_imm =>
+ returnm (true : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
: M (bool)))
: M (bool))
- (returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__7 14 13)
- (vec_of_bits [B1;B1]
- : mword (14 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1]
- : mword (6 - 0 + 1)))))
- : bool))) >>= fun w__56 : bool =>
- (if sumbool_of_bool (w__56) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_41_ : mword 1 :=
- subrange_vec_dec v__7 12 12 in
- (bool_not_bits_backwards _mappingpatterns_41_) >>= fun s =>
- returnm ((REM
- ((rs2, rs1, rd, s)))
- : ast)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec
- (subrange_vec_dec v__7
- 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec
- (subrange_vec_dec
- v__7 14 12)
- (vec_of_bits [B0;B0;B0]
- : mword (14 - 12 + 1)))
- (eq_vec
- (subrange_vec_dec
- v__7 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1)))))))
- then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- returnm ((MULW
- ((rs2, rs1, rd)))
- : ast)
+ (returnm ((eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"1110011"
+ : mword (6 - 0 + 1)))
+ : bool))) >>= fun w__58 : bool =>
+ (if sumbool_of_bool w__58 then
+ let csr : mword 12 := subrange_vec_dec v__9 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let csr : mword 12 := subrange_vec_dec v__9 31 20 in
+ let _mappingpatterns_81_ : mword 2 :=
+ subrange_vec_dec v__9 13 12 in
+ let _mappingpatterns_80_ : mword 1 :=
+ subrange_vec_dec v__9 14 14 in
+ (encdec_csrop_backwards _mappingpatterns_81_) >>= fun op =>
+ (bool_bits_backwards _mappingpatterns_80_) >>= fun is_imm =>
+ returnm (CSR (csr, rs1, rd, is_imm, op))
+ else if eq_vec v__9 (Ox"00200073" : mword 32) then
+ returnm (URET tt)
else
- (and_boolM
- (let _mappingpatterns_42_ : mword 1 :=
- subrange_vec_dec v__7 12 12 in
- (and_boolM
- (returnm ((bool_not_bits_backwards_matches
- _mappingpatterns_42_)
- : bool))
- ((if ((bool_not_bits_backwards_matches
- _mappingpatterns_42_)) then
- (bool_not_bits_backwards _mappingpatterns_42_) >>= fun s =>
- returnm (((Z.eqb 64 64)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool)))
- : M (bool))
+ (and_boolM ((is_RV32F_or_RV64F tt) : M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec
- (subrange_vec_dec v__7 14 13)
- (vec_of_bits [B1;B0]
- : mword (14 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1)))))
- : bool))) >>= fun w__59 : bool =>
- (if sumbool_of_bool (w__59) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_42_ : mword 1 :=
- subrange_vec_dec v__7 12 12 in
- (bool_not_bits_backwards _mappingpatterns_42_) >>= fun s =>
- returnm ((DIVW
- ((rs2, rs1, rd, s)))
- : ast)
+ (eq_vec (subrange_vec_dec v__9 14 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0000111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__60 : bool =>
+ (if sumbool_of_bool w__60 then
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ returnm (LOAD_FP (imm, rs1, rd, WORD))
else
- (and_boolM
- (let _mappingpatterns_43_ : mword 1 :=
- subrange_vec_dec v__7 12 12 in
- (and_boolM
- (returnm ((bool_not_bits_backwards_matches
- _mappingpatterns_43_)
- : bool))
- ((if ((bool_not_bits_backwards_matches
- _mappingpatterns_43_)) then
- (bool_not_bits_backwards
- _mappingpatterns_43_) >>= fun s =>
- returnm (((Z.eqb 64 64)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool)))
- : M (bool))
+ (and_boolM ((is_RV32D_or_RV64D tt) : M (bool))
(returnm ((andb
(eq_vec
- (subrange_vec_dec v__7 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec
- (subrange_vec_dec v__7 14 13)
- (vec_of_bits [B1;B1]
- : mword (14 - 13 + 1)))
- (eq_vec
- (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1)))))
+ (subrange_vec_dec v__9 14 12)
+ ('b"011"
+ : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__9 6 0)
+ ('b"0000111"
+ : mword (6 - 0 + 1))))
: bool))) >>= fun w__62 : bool =>
- (if sumbool_of_bool (w__62) then
- let rs2 : mword 5 := subrange_vec_dec v__7 24 20 in
- let rs1 : mword 5 := subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let _mappingpatterns_43_ : mword 1 :=
- subrange_vec_dec v__7 12 12 in
- (bool_not_bits_backwards _mappingpatterns_43_) >>= fun s =>
- returnm ((REMW
- ((rs2, rs1, rd, s)))
- : ast)
+ (if sumbool_of_bool w__62 then
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ let rs1 : mword 5 := subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 := subrange_vec_dec v__9 11 7 in
+ let imm : mword 12 := subrange_vec_dec v__9 31 20 in
+ returnm (LOAD_FP (imm, rs1, rd, DOUBLE))
else
- (and_boolM
- (let _mappingpatterns_45_ : mword 2 :=
- subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_44_ : mword 1 :=
- subrange_vec_dec v__7 14 14 in
- (and_boolM
- (returnm ((encdec_csrop_backwards_matches
- _mappingpatterns_45_)
- : bool))
- ((if ((encdec_csrop_backwards_matches
- _mappingpatterns_45_)) then
- (encdec_csrop_backwards
- _mappingpatterns_45_) >>= fun op =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_44_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_44_)) then
- (bool_bits_backwards
- _mappingpatterns_44_) >>= fun is_imm =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool))) >>= fun w__64 : bool =>
- returnm ((w__64
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool)))
- : M (bool))
- (returnm ((eq_vec (subrange_vec_dec v__7 6 0)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1]
- : mword (6 - 0 + 1)))
- : bool))) >>= fun w__67 : bool =>
- (if sumbool_of_bool (w__67) then
- let csr : mword 12 :=
- subrange_vec_dec v__7 31 20 in
+ (and_boolM ((is_RV32F_or_RV64F tt) : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec v__9 14 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec v__9 6 0)
+ ('b"0100111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__64 : bool =>
+ (if sumbool_of_bool w__64 then
+ let imm7 : bits 7 :=
+ subrange_vec_dec v__9 31 25 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec v__9 24 20 in
let rs1 : mword 5 :=
- subrange_vec_dec v__7 19 15 in
- let rd : mword 5 := subrange_vec_dec v__7 11 7 in
- let csr : mword 12 :=
- subrange_vec_dec v__7 31 20 in
- let _mappingpatterns_45_ : mword 2 :=
- subrange_vec_dec v__7 13 12 in
- let _mappingpatterns_44_ : mword 1 :=
- subrange_vec_dec v__7 14 14 in
- (encdec_csrop_backwards _mappingpatterns_45_) >>= fun op =>
- (bool_bits_backwards _mappingpatterns_44_) >>= fun is_imm =>
- returnm ((CSR
- ((csr, rs1, rd, is_imm, op)))
- : ast)
+ subrange_vec_dec v__9 19 15 in
+ let imm7 : bits 7 :=
+ subrange_vec_dec v__9 31 25 in
+ let imm5 : bits 5 :=
+ subrange_vec_dec v__9 11 7 in
+ returnm (STORE_FP
+ (concat_vec (imm7 : bits 7)
+ (imm5
+ : bits 5), rs2, rs1, WORD))
else
- returnm ((if ((eq_vec v__7
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- URET
- (tt)
- else ILLEGAL (v__7))
- : ast))
+ (and_boolM
+ ((is_RV32D_or_RV64D tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec v__9 14
+ 12)
+ ('b"011"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec v__9 6 0)
+ ('b"0100111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__66 : bool =>
+ (if sumbool_of_bool w__66 then
+ let imm7 : bits 7 :=
+ subrange_vec_dec v__9 31 25 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec v__9 19 15 in
+ let imm7 : bits 7 :=
+ subrange_vec_dec v__9 31 25 in
+ let imm5 : bits 5 :=
+ subrange_vec_dec v__9 11 7 in
+ returnm (STORE_FP
+ (concat_vec (imm7 : bits 7)
+ (imm5
+ : bits 5), rs2, rs1, DOUBLE))
+ else
+ (and_boolM
+ (let _mappingpatterns_82_ : mword 3 :=
+ subrange_vec_dec v__9 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_82_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_82_ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_82_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec v__9
+ 26 25)
+ ('b"00"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec v__9 6
+ 0)
+ ('b"1000011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__70 : bool =>
+ (if sumbool_of_bool w__70 then
+ let rs3 : mword 5 :=
+ subrange_vec_dec v__9 31 27 in
+ let rs3 : mword 5 :=
+ subrange_vec_dec v__9 31 27 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_82_ : mword 3 :=
+ subrange_vec_dec v__9 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_82_) >>= fun rm =>
+ returnm (F_MADD_TYPE_S
+ (rs3, rs2, rs1, rm, rd, FMADD_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_83_ : mword 3 :=
+ subrange_vec_dec v__9 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_83_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_83_ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_83_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9 26 25)
+ ('b"00"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9 6 0)
+ ('b"1000111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__74 : bool =>
+ (if sumbool_of_bool w__74 then
+ let rs3 : mword 5 :=
+ subrange_vec_dec v__9 31 27 in
+ let rs3 : mword 5 :=
+ subrange_vec_dec v__9 31 27 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_83_ : mword 3 :=
+ subrange_vec_dec v__9 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_83_) >>= fun rm =>
+ returnm (F_MADD_TYPE_S
+ (rs3, rs2, rs1, rm, rd, FMSUB_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_84_ : mword 3 :=
+ subrange_vec_dec v__9 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_84_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_84_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_84_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9 26 25)
+ ('b"00"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9 6 0)
+ ('b"1001011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__78 : bool =>
+ (if sumbool_of_bool w__78 then
+ let rs3 : mword 5 :=
+ subrange_vec_dec v__9 31 27 in
+ let rs3 : mword 5 :=
+ subrange_vec_dec v__9 31 27 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_84_ : mword 3 :=
+ subrange_vec_dec v__9 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_84_) >>= fun rm =>
+ returnm (F_MADD_TYPE_S
+ (rs3, rs2, rs1, rm, rd, FNMSUB_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_85_ : mword 3 :=
+ subrange_vec_dec v__9 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_85_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_85_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_85_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9 26 25)
+ ('b"00"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9 6 0)
+ ('b"1001111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__82 : bool =>
+ (if sumbool_of_bool w__82 then
+ let rs3 : mword 5 :=
+ subrange_vec_dec v__9 31 27 in
+ let rs3 : mword 5 :=
+ subrange_vec_dec v__9 31 27 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec v__9 24 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec v__9 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_85_ : mword 3 :=
+ subrange_vec_dec v__9 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_85_) >>= fun rm =>
+ returnm (F_MADD_TYPE_S
+ (rs3, rs2, rs1, rm, rd, FNMADD_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_86_ : mword 3 :=
+ subrange_vec_dec v__9 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_86_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_86_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_86_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9 31
+ 25)
+ ('b"0000000"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9 6 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__86 : bool =>
+ (if sumbool_of_bool w__86 then
+ let rs2 : mword 5 :=
+ subrange_vec_dec v__9 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec v__9 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec v__9 11 7 in
+ let _mappingpatterns_86_ : mword 3 :=
+ subrange_vec_dec v__9 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_86_) >>= fun rm =>
+ returnm (F_BIN_RM_TYPE_S
+ (rs2, rs1, rm, rd, FADD_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_87_ : mword 3 :=
+ subrange_vec_dec v__9
+ 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_87_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_87_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_87_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31 25)
+ ('b"0000100"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__90 : bool =>
+ (if sumbool_of_bool
+ w__90 then
+ let rs2 : mword 5 :=
+ subrange_vec_dec v__9
+ 24 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec v__9
+ 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec v__9
+ 11 7 in
+ let _mappingpatterns_87_ : mword 3 :=
+ subrange_vec_dec v__9
+ 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_87_) >>= fun rm =>
+ returnm (F_BIN_RM_TYPE_S
+ (rs2, rs1, rm, rd, FSUB_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_88_ : mword 3 :=
+ subrange_vec_dec
+ v__9 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_88_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_88_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_88_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0001000"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__94 : bool =>
+ (if sumbool_of_bool
+ w__94 then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9 24 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9 11 7 in
+ let _mappingpatterns_88_ : mword 3 :=
+ subrange_vec_dec
+ v__9 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_88_) >>= fun rm =>
+ returnm (F_BIN_RM_TYPE_S
+ (rs2, rs1, rm, rd, FMUL_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_89_ : mword 3 :=
+ subrange_vec_dec
+ v__9 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_89_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_89_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_89_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0001100"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__98 : bool =>
+ (if sumbool_of_bool
+ w__98 then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9 24 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9 11 7 in
+ let _mappingpatterns_89_ : mword 3 :=
+ subrange_vec_dec
+ v__9 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_89_) >>= fun rm =>
+ returnm (F_BIN_RM_TYPE_S
+ (rs2, rs1, rm, rd, FDIV_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_90_ : mword 3 :=
+ subrange_vec_dec
+ v__9 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_90_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_90_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_90_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"580"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__102 : bool =>
+ (if sumbool_of_bool
+ w__102 then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9 11 7 in
+ let _mappingpatterns_90_ : mword 3 :=
+ subrange_vec_dec
+ v__9 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_90_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FSQRT_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_91_ : mword 3 :=
+ subrange_vec_dec
+ v__9 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_91_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_91_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_91_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"C00"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__106 : bool =>
+ (if sumbool_of_bool
+ w__106 then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9 11 7 in
+ let _mappingpatterns_91_ : mword 3 :=
+ subrange_vec_dec
+ v__9 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_91_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_W_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_92_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_92_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_92_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_92_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"C01"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__110 : bool =>
+ (if sumbool_of_bool
+ w__110
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11 7 in
+ let _mappingpatterns_92_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_92_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_WU_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_93_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_93_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_93_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_93_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"D00"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__114 : bool =>
+ (if sumbool_of_bool
+ w__114
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_93_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_93_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_W))
+ else
+ (and_boolM
+ (let _mappingpatterns_94_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_94_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_94_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_94_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"D01"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__118 : bool =>
+ (if sumbool_of_bool
+ w__118
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_94_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_94_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_WU))
+ else
+ (and_boolM
+ (let _mappingpatterns_95_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_95_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_95_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_95_) >>= fun rm =>
+ (is_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"C02"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__122 : bool =>
+ (if sumbool_of_bool
+ w__122
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_95_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_95_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_L_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_96_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_96_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_96_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_96_) >>= fun rm =>
+ (is_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"C03"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__126 : bool =>
+ (if
+ sumbool_of_bool
+ w__126
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_96_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_96_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_LU_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_97_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_97_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_97_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_97_) >>= fun rm =>
+ (is_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"D02"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__130 : bool =>
+ (if
+ sumbool_of_bool
+ w__130
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_97_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_97_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_L))
+ else
+ (and_boolM
+ (let _mappingpatterns_98_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_98_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_98_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_98_) >>= fun rm =>
+ (is_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"D03"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__134 : bool =>
+ (if
+ sumbool_of_bool
+ w__134
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_98_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_98_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_LU))
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__136 : bool =>
+ (if
+ sumbool_of_bool
+ w__136
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJ_S))
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__138 : bool =>
+ (if
+ sumbool_of_bool
+ w__138
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJN_S))
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__140 : bool =>
+ (if
+ sumbool_of_bool
+ w__140
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJX_S))
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010100"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__142 : bool =>
+ (if
+ sumbool_of_bool
+ w__142
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FMIN_S))
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010100"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__144 : bool =>
+ (if
+ sumbool_of_bool
+ w__144
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FMAX_S))
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"1010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__146 : bool =>
+ (if
+ sumbool_of_bool
+ w__146
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FEQ_S))
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"1010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__148 : bool =>
+ (if
+ sumbool_of_bool
+ w__148
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FLT_S))
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"1010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__150 : bool =>
+ (if
+ sumbool_of_bool
+ w__150
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FLE_S))
+ else
+ (and_boolM
+ ((haveFExt
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"E00"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__152 : bool =>
+ (if
+ sumbool_of_bool
+ w__152
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_UN_TYPE_S
+ (rs1, rd, FCLASS_S))
+ else
+ (and_boolM
+ ((haveFExt
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"E00"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__154 : bool =>
+ (if
+ sumbool_of_bool
+ w__154
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_UN_TYPE_S
+ (rs1, rd, FMV_X_W))
+ else
+ (and_boolM
+ ((haveFExt
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"F00"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__156 : bool =>
+ (if
+ sumbool_of_bool
+ w__156
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_UN_TYPE_S
+ (rs1, rd, FMV_W_X))
+ else
+ (and_boolM
+ (let _mappingpatterns_99_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_99_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_99_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_99_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 26
+ 25)
+ ('b"01"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1000011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__160 : bool =>
+ (if
+ sumbool_of_bool
+ w__160
+ then
+ let rs3 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 31
+ 27 in
+ let rs3 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 31
+ 27 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_99_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_99_) >>= fun rm =>
+ returnm (F_MADD_TYPE_D
+ (rs3, rs2, rs1, rm, rd, FMADD_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_100_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_100_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_100_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_100_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 26
+ 25)
+ ('b"01"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1000111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__164 : bool =>
+ (if
+ sumbool_of_bool
+ w__164
+ then
+ let rs3 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 31
+ 27 in
+ let rs3 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 31
+ 27 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_100_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_100_) >>= fun rm =>
+ returnm (F_MADD_TYPE_D
+ (rs3, rs2, rs1, rm, rd, FMSUB_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_101_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_101_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_101_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_101_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 26
+ 25)
+ ('b"01"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1001011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__168 : bool =>
+ (if
+ sumbool_of_bool
+ w__168
+ then
+ let rs3 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 31
+ 27 in
+ let rs3 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 31
+ 27 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_101_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_101_) >>= fun rm =>
+ returnm (F_MADD_TYPE_D
+ (rs3, rs2, rs1, rm, rd, FNMSUB_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_102_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_102_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_102_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_102_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 26
+ 25)
+ ('b"01"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1001111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__172 : bool =>
+ (if
+ sumbool_of_bool
+ w__172
+ then
+ let rs3 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 31
+ 27 in
+ let rs3 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 31
+ 27 in
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_102_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_102_) >>= fun rm =>
+ returnm (F_MADD_TYPE_D
+ (rs3, rs2, rs1, rm, rd, FNMADD_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_103_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_103_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_103_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_103_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__176 : bool =>
+ (if
+ sumbool_of_bool
+ w__176
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_103_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_103_) >>= fun rm =>
+ returnm (F_BIN_RM_TYPE_D
+ (rs2, rs1, rm, rd, FADD_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_104_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_104_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_104_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_104_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0000101"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__180 : bool =>
+ (if
+ sumbool_of_bool
+ w__180
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_104_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_104_) >>= fun rm =>
+ returnm (F_BIN_RM_TYPE_D
+ (rs2, rs1, rm, rd, FSUB_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_105_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_105_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_105_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_105_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0001001"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__184 : bool =>
+ (if
+ sumbool_of_bool
+ w__184
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_105_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_105_) >>= fun rm =>
+ returnm (F_BIN_RM_TYPE_D
+ (rs2, rs1, rm, rd, FMUL_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_106_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_106_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_106_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_106_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0001101"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__188 : bool =>
+ (if
+ sumbool_of_bool
+ w__188
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_106_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_106_) >>= fun rm =>
+ returnm (F_BIN_RM_TYPE_D
+ (rs2, rs1, rm, rd, FDIV_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_107_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_107_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_107_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_107_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"5A0"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__192 : bool =>
+ (if
+ sumbool_of_bool
+ w__192
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_107_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_107_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FSQRT_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_108_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_108_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_108_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_108_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"C20"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__196 : bool =>
+ (if
+ sumbool_of_bool
+ w__196
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_108_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_108_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_W_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_109_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_109_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_109_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_109_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"C21"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__200 : bool =>
+ (if
+ sumbool_of_bool
+ w__200
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_109_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_109_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_WU_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_110_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_110_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_110_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_110_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"D20"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__204 : bool =>
+ (if
+ sumbool_of_bool
+ w__204
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_110_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_110_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_W))
+ else
+ (and_boolM
+ (let _mappingpatterns_111_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_111_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_111_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_111_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"D21"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__208 : bool =>
+ (if
+ sumbool_of_bool
+ w__208
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_111_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_111_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_WU))
+ else
+ (and_boolM
+ (let _mappingpatterns_112_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_112_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_112_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_112_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"401"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__212 : bool =>
+ (if
+ sumbool_of_bool
+ w__212
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_112_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_112_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_S_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_113_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_113_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_113_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_113_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"420"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__216 : bool =>
+ (if
+ sumbool_of_bool
+ w__216
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_113_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_113_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_S))
+ else
+ (and_boolM
+ (let _mappingpatterns_114_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_114_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_114_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_114_) >>= fun rm =>
+ (is_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"C22"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__220 : bool =>
+ (if
+ sumbool_of_bool
+ w__220
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_114_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_114_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_L_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_115_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_115_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_115_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_115_) >>= fun rm =>
+ (is_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"C23"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__224 : bool =>
+ (if
+ sumbool_of_bool
+ w__224
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_115_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_115_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_LU_D))
+ else
+ (and_boolM
+ (let _mappingpatterns_116_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_116_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_116_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_116_) >>= fun rm =>
+ (is_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"D22"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__228 : bool =>
+ (if
+ sumbool_of_bool
+ w__228
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_116_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_116_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_L))
+ else
+ (and_boolM
+ (let _mappingpatterns_117_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_117_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_117_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_117_) >>= fun rm =>
+ (is_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"D23"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__232 : bool =>
+ (if
+ sumbool_of_bool
+ w__232
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ let _mappingpatterns_117_ : mword 3 :=
+ subrange_vec_dec
+ v__9
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_117_) >>= fun rm =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_LU))
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__234 : bool =>
+ (if
+ sumbool_of_bool
+ w__234
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJ_D))
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__236 : bool =>
+ (if
+ sumbool_of_bool
+ w__236
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJN_D))
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__238 : bool =>
+ (if
+ sumbool_of_bool
+ w__238
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJX_D))
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010101"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__240 : bool =>
+ (if
+ sumbool_of_bool
+ w__240
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FMIN_D))
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"0010101"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__242 : bool =>
+ (if
+ sumbool_of_bool
+ w__242
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FMAX_D))
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"1010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__244 : bool =>
+ (if
+ sumbool_of_bool
+ w__244
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FEQ_D))
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"1010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__246 : bool =>
+ (if
+ sumbool_of_bool
+ w__246
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FLT_D))
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 25)
+ ('b"1010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__248 : bool =>
+ (if
+ sumbool_of_bool
+ w__248
+ then
+ let rs2 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 24
+ 20 in
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FLE_D))
+ else
+ (and_boolM
+ ((haveDExt
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"E20"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__250 : bool =>
+ (if
+ sumbool_of_bool
+ w__250
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_UN_TYPE_D
+ (rs1, rd, FCLASS_D))
+ else
+ (and_boolM
+ ((is_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"E20"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__252 : bool =>
+ (if
+ sumbool_of_bool
+ w__252
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ returnm (F_UN_TYPE_D
+ (rs1, rd, FMV_X_D))
+ else
+ (and_boolM
+ ((is_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 31
+ 20)
+ (Ox"F20"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__9
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__254 : bool =>
+ returnm (if
+ sumbool_of_bool
+ w__254
+ then
+ let rs1 : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 19
+ 15 in
+ let rd : mword 5 :=
+ subrange_vec_dec
+ v__9
+ 11
+ 7 in
+ F_UN_TYPE_D
+ (rs1, rd, FMV_D_X)
+ else
+ ILLEGAL
+ v__9))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
: M (ast))
: M (ast))
: M (ast))
@@ -19128,1280 +23280,3541 @@ Definition encdec_backwards (arg_ : mword 32)
: M (ast))
: M (ast).
-Definition encdec_forwards_matches (arg_ : ast)
-: bool :=
-
- match arg_ with
- | UTYPE ((imm, rd, op)) => true
- | RISCV_JAL ((v__220, rd)) =>
- if ((eq_vec (subrange_vec_dec v__220 0 0) (vec_of_bits [B0] : mword (0 - 0 + 1)))) then true
- else false
- | RISCV_JALR ((imm, rs1, rd)) => true
- | BTYPE ((v__222, rs2, rs1, op)) =>
- if ((eq_vec (subrange_vec_dec v__222 0 0) (vec_of_bits [B0] : mword (0 - 0 + 1)))) then true
- else false
- | ITYPE ((imm, rs1, rd, op)) => true
- | SHIFTIOP ((shamt, rs1, rd, RISCV_SLLI)) => true
- | SHIFTIOP ((shamt, rs1, rd, RISCV_SRLI)) => true
- | SHIFTIOP ((shamt, rs1, rd, RISCV_SRAI)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_ADD)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_SLT)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_SLTU)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_AND)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_OR)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_XOR)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_SLL)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_SRL)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_SUB)) => true
- | RTYPE ((rs2, rs1, rd, RISCV_SRA)) => true
- | LOAD ((imm, rs1, rd, is_unsigned, size, false, false)) =>
- if sumbool_of_bool ((orb (Z.ltb (projT1 (word_width_bytes size)) 8)
- (andb (negb is_unsigned)
- ((Z.leb (projT1 (word_width_bytes size)) 8)
- : bool)))) then
- true
- else false
- | STORE ((v__224, rs2, rs1, size, false, false)) =>
- if sumbool_of_bool ((Z.leb (projT1 (word_width_bytes size)) 8)) then true else false
- | ADDIW ((imm, rs1, rd)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | SHIFTW ((shamt, rs1, rd, RISCV_SLLI)) =>
- if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | SHIFTW ((shamt, rs1, rd, RISCV_SRLI)) =>
- if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | SHIFTW ((shamt, rs1, rd, RISCV_SRAI)) =>
- if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | RTYPEW ((rs2, rs1, rd, RISCV_ADDW)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | RTYPEW ((rs2, rs1, rd, RISCV_SUBW)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | RTYPEW ((rs2, rs1, rd, RISCV_SLLW)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | RTYPEW ((rs2, rs1, rd, RISCV_SRLW)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | RTYPEW ((rs2, rs1, rd, RISCV_SRAW)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | SHIFTIWOP ((shamt, rs1, rd, RISCV_SLLIW)) =>
- if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | SHIFTIWOP ((shamt, rs1, rd, RISCV_SRLIW)) =>
- if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | SHIFTIWOP ((shamt, rs1, rd, RISCV_SRAIW)) =>
- if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | FENCE ((pred, succ)) => true
- | FENCE_TSO ((pred, succ)) => true
- | FENCEI (tt) => true
- | ECALL (tt) => true
- | MRET (tt) => true
- | SRET (tt) => true
- | EBREAK (tt) => true
- | WFI (tt) => true
- | SFENCE_VMA ((rs1, rs2)) => true
- | LOADRES ((aq, rl, rs1, size, rd)) =>
- if sumbool_of_bool ((Z.leb (projT1 (word_width_bytes size)) 8)) then true else false
- | STORECON ((aq, rl, rs2, rs1, size, rd)) =>
- if sumbool_of_bool ((Z.leb (projT1 (word_width_bytes size)) 8)) then true else false
- | AMO ((op, aq, rl, rs2, rs1, size, rd)) =>
- if sumbool_of_bool ((Z.leb (projT1 (word_width_bytes size)) 8)) then true else false
- | MUL ((rs2, rs1, rd, high, signed1, signed2)) => true
- | DIV ((rs2, rs1, rd, s)) => true
- | REM ((rs2, rs1, rd, s)) => true
- | MULW ((rs2, rs1, rd)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | DIVW ((rs2, rs1, rd, s)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | REMW ((rs2, rs1, rd, s)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | CSR ((csr, rs1, rd, is_imm, op)) => true
- | URET (tt) => true
- | ILLEGAL (s) => true
- | _ => false
- end.
+Definition encdec_forwards_matches (arg_ : ast) : M (bool) :=
+ (match arg_ with
+ | UTYPE (imm, rd, op) => returnm true
+ | RISCV_JAL (v__456, rd) =>
+ returnm (if eq_vec (subrange_vec_dec v__456 0 0) ('b"0" : mword (0 - 0 + 1)) then true
+ else false)
+ | RISCV_JALR (imm, rs1, rd) => returnm true
+ | BTYPE (v__458, rs2, rs1, op) =>
+ returnm (if eq_vec (subrange_vec_dec v__458 0 0) ('b"0" : mword (0 - 0 + 1)) then true
+ else false)
+ | ITYPE (imm, rs1, rd, op) => returnm true
+ | SHIFTIOP (shamt, rs1, rd, RISCV_SLLI) => returnm true
+ | SHIFTIOP (shamt, rs1, rd, RISCV_SRLI) => returnm true
+ | SHIFTIOP (shamt, rs1, rd, RISCV_SRAI) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_ADD) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_SLT) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_SLTU) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_AND) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_OR) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_XOR) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_SLL) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_SRL) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_SUB) => returnm true
+ | RTYPE (rs2, rs1, rd, RISCV_SRA) => returnm true
+ | LOAD (imm, rs1, rd, is_unsigned, size, false, false) =>
+ returnm (if sumbool_of_bool
+ (orb (Z.ltb (projT1 (word_width_bytes size)) 8)
+ (andb (negb is_unsigned) ((Z.leb (projT1 (word_width_bytes size)) 8) : bool)))
+ then
+ true
+ else false)
+ | STORE (v__460, rs2, rs1, size, false, false) =>
+ returnm (if sumbool_of_bool (Z.leb (projT1 (word_width_bytes size)) 8) then true else false)
+ | ADDIW (imm, rs1, rd) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | SHIFTW (shamt, rs1, rd, RISCV_SLLI) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | SHIFTW (shamt, rs1, rd, RISCV_SRLI) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | SHIFTW (shamt, rs1, rd, RISCV_SRAI) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | RTYPEW (rs2, rs1, rd, RISCV_ADDW) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | RTYPEW (rs2, rs1, rd, RISCV_SUBW) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | RTYPEW (rs2, rs1, rd, RISCV_SLLW) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | RTYPEW (rs2, rs1, rd, RISCV_SRLW) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | RTYPEW (rs2, rs1, rd, RISCV_SRAW) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | SHIFTIWOP (shamt, rs1, rd, RISCV_SLLIW) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | SHIFTIWOP (shamt, rs1, rd, RISCV_SRLIW) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | SHIFTIWOP (shamt, rs1, rd, RISCV_SRAIW) =>
+ returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | FENCE (pred, succ) => returnm true
+ | FENCE_TSO (pred, succ) => returnm true
+ | FENCEI tt => returnm true
+ | ECALL tt => returnm true
+ | MRET tt => returnm true
+ | SRET tt => returnm true
+ | EBREAK tt => returnm true
+ | WFI tt => returnm true
+ | SFENCE_VMA (rs1, rs2) => returnm true
+ | LOADRES (aq, rl, rs1, size, rd) =>
+ returnm (if sumbool_of_bool (Z.leb (projT1 (word_width_bytes size)) 8) then true else false)
+ | STORECON (aq, rl, rs2, rs1, size, rd) =>
+ returnm (if sumbool_of_bool (Z.leb (projT1 (word_width_bytes size)) 8) then true else false)
+ | AMO (op, aq, rl, rs2, rs1, size, rd) =>
+ returnm (if sumbool_of_bool (Z.leb (projT1 (word_width_bytes size)) 8) then true else false)
+ | MUL (rs2, rs1, rd, high, signed1, signed2) => returnm true
+ | DIV (rs2, rs1, rd, s) => returnm true
+ | REM (rs2, rs1, rd, s) => returnm true
+ | MULW (rs2, rs1, rd) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | DIVW (rs2, rs1, rd, s) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | REMW (rs2, rs1, rd, s) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | CSR (csr, rs1, rd, is_imm, op) => returnm true
+ | URET tt => returnm true
+ | LOAD_FP (imm, rs1, rd, WORD) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__0 : bool =>
+ returnm (if sumbool_of_bool w__0 then true else false)
+ | LOAD_FP (imm, rs1, rd, DOUBLE) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__1 : bool =>
+ returnm (if sumbool_of_bool w__1 then true else false)
+ | STORE_FP (v__461, rs2, rs1, WORD) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__2 : bool =>
+ returnm (if sumbool_of_bool w__2 then true else false)
+ | STORE_FP (v__462, rs2, rs1, DOUBLE) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__3 : bool =>
+ returnm (if sumbool_of_bool w__3 then true else false)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FMADD_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__4 : bool =>
+ returnm (if sumbool_of_bool w__4 then true else false)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FMSUB_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__5 : bool =>
+ returnm (if sumbool_of_bool w__5 then true else false)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FNMSUB_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__6 : bool =>
+ returnm (if sumbool_of_bool w__6 then true else false)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FNMADD_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__7 : bool =>
+ returnm (if sumbool_of_bool w__7 then true else false)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FADD_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__8 : bool =>
+ returnm (if sumbool_of_bool w__8 then true else false)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FSUB_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__9 : bool =>
+ returnm (if sumbool_of_bool w__9 then true else false)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FMUL_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__10 : bool =>
+ returnm (if sumbool_of_bool w__10 then true else false)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FDIV_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__11 : bool =>
+ returnm (if sumbool_of_bool w__11 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FSQRT_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__12 : bool =>
+ returnm (if sumbool_of_bool w__12 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_W_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__13 : bool =>
+ returnm (if sumbool_of_bool w__13 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_WU_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__14 : bool =>
+ returnm (if sumbool_of_bool w__14 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_W) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__15 : bool =>
+ returnm (if sumbool_of_bool w__15 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_WU) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__16 : bool =>
+ returnm (if sumbool_of_bool w__16 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_L_S) =>
+ (is_RV64F tt) >>= fun w__17 : bool => returnm (if sumbool_of_bool w__17 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_LU_S) =>
+ (is_RV64F tt) >>= fun w__18 : bool => returnm (if sumbool_of_bool w__18 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_L) =>
+ (is_RV64F tt) >>= fun w__19 : bool => returnm (if sumbool_of_bool w__19 then true else false)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_LU) =>
+ (is_RV64F tt) >>= fun w__20 : bool => returnm (if sumbool_of_bool w__20 then true else false)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJ_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__21 : bool =>
+ returnm (if sumbool_of_bool w__21 then true else false)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJN_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__22 : bool =>
+ returnm (if sumbool_of_bool w__22 then true else false)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJX_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__23 : bool =>
+ returnm (if sumbool_of_bool w__23 then true else false)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FMIN_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__24 : bool =>
+ returnm (if sumbool_of_bool w__24 then true else false)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FMAX_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__25 : bool =>
+ returnm (if sumbool_of_bool w__25 then true else false)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FEQ_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__26 : bool =>
+ returnm (if sumbool_of_bool w__26 then true else false)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FLT_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__27 : bool =>
+ returnm (if sumbool_of_bool w__27 then true else false)
+ | F_BIN_TYPE_S (rs2, rs1, rd, FLE_S) =>
+ (is_RV32F_or_RV64F tt) >>= fun w__28 : bool =>
+ returnm (if sumbool_of_bool w__28 then true else false)
+ | F_UN_TYPE_S (rs1, rd, FCLASS_S) =>
+ (haveFExt tt) >>= fun w__29 : bool => returnm (if sumbool_of_bool w__29 then true else false)
+ | F_UN_TYPE_S (rs1, rd, FMV_X_W) =>
+ (haveFExt tt) >>= fun w__30 : bool => returnm (if sumbool_of_bool w__30 then true else false)
+ | F_UN_TYPE_S (rs1, rd, FMV_W_X) =>
+ (haveFExt tt) >>= fun w__31 : bool => returnm (if sumbool_of_bool w__31 then true else false)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, FMADD_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__32 : bool =>
+ returnm (if sumbool_of_bool w__32 then true else false)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, FMSUB_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__33 : bool =>
+ returnm (if sumbool_of_bool w__33 then true else false)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, FNMSUB_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__34 : bool =>
+ returnm (if sumbool_of_bool w__34 then true else false)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, FNMADD_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__35 : bool =>
+ returnm (if sumbool_of_bool w__35 then true else false)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, FADD_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__36 : bool =>
+ returnm (if sumbool_of_bool w__36 then true else false)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, FSUB_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__37 : bool =>
+ returnm (if sumbool_of_bool w__37 then true else false)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, FMUL_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__38 : bool =>
+ returnm (if sumbool_of_bool w__38 then true else false)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, FDIV_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__39 : bool =>
+ returnm (if sumbool_of_bool w__39 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FSQRT_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__40 : bool =>
+ returnm (if sumbool_of_bool w__40 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_W_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__41 : bool =>
+ returnm (if sumbool_of_bool w__41 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_WU_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__42 : bool =>
+ returnm (if sumbool_of_bool w__42 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_W) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__43 : bool =>
+ returnm (if sumbool_of_bool w__43 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_WU) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__44 : bool =>
+ returnm (if sumbool_of_bool w__44 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_S_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__45 : bool =>
+ returnm (if sumbool_of_bool w__45 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_S) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__46 : bool =>
+ returnm (if sumbool_of_bool w__46 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_L_D) =>
+ (is_RV64D tt) >>= fun w__47 : bool => returnm (if sumbool_of_bool w__47 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_LU_D) =>
+ (is_RV64D tt) >>= fun w__48 : bool => returnm (if sumbool_of_bool w__48 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_L) =>
+ (is_RV64D tt) >>= fun w__49 : bool => returnm (if sumbool_of_bool w__49 then true else false)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_LU) =>
+ (is_RV64D tt) >>= fun w__50 : bool => returnm (if sumbool_of_bool w__50 then true else false)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJ_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__51 : bool =>
+ returnm (if sumbool_of_bool w__51 then true else false)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJN_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__52 : bool =>
+ returnm (if sumbool_of_bool w__52 then true else false)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJX_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__53 : bool =>
+ returnm (if sumbool_of_bool w__53 then true else false)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FMIN_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__54 : bool =>
+ returnm (if sumbool_of_bool w__54 then true else false)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FMAX_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__55 : bool =>
+ returnm (if sumbool_of_bool w__55 then true else false)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FEQ_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__56 : bool =>
+ returnm (if sumbool_of_bool w__56 then true else false)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FLT_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__57 : bool =>
+ returnm (if sumbool_of_bool w__57 then true else false)
+ | F_BIN_TYPE_D (rs2, rs1, rd, FLE_D) =>
+ (is_RV32D_or_RV64D tt) >>= fun w__58 : bool =>
+ returnm (if sumbool_of_bool w__58 then true else false)
+ | F_UN_TYPE_D (rs1, rd, FCLASS_D) =>
+ (haveDExt tt) >>= fun w__59 : bool => returnm (if sumbool_of_bool w__59 then true else false)
+ | F_UN_TYPE_D (rs1, rd, FMV_X_D) =>
+ (is_RV64D tt) >>= fun w__60 : bool => returnm (if sumbool_of_bool w__60 then true else false)
+ | F_UN_TYPE_D (rs1, rd, FMV_D_X) =>
+ (is_RV64D tt) >>= fun w__61 : bool => returnm (if sumbool_of_bool w__61 then true else false)
+ | ILLEGAL s => returnm true
+ | _ => returnm false
+ end)
+ : M (bool).
-Definition encdec_backwards_matches (arg_ : mword 32)
-: M (bool) :=
-
- let v__225 := arg_ in
- let _mappingpatterns_0_ : mword 7 := subrange_vec_dec v__225 6 0 in
+Definition encdec_backwards_matches (arg_ : mword 32) : M (bool) :=
+ let v__463 := arg_ in
+ let _mappingpatterns_0_ : mword 7 := subrange_vec_dec v__463 6 0 in
(and_boolM (returnm ((encdec_uop_backwards_matches _mappingpatterns_0_) : bool))
- ((if ((encdec_uop_backwards_matches _mappingpatterns_0_)) then
- (encdec_uop_backwards _mappingpatterns_0_) >>= fun op => returnm ((true : bool) : bool)
- else returnm (projT1 (build_ex false : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if encdec_uop_backwards_matches _mappingpatterns_0_ then
+ (encdec_uop_backwards _mappingpatterns_0_) >>= fun op => returnm (true : bool)
+ else returnm false)
: M (bool))) >>= fun w__1 : bool =>
- (if sumbool_of_bool (w__1) then
- let _mappingpatterns_0_ : mword 7 := subrange_vec_dec v__225 6 0 in
- (encdec_uop_backwards _mappingpatterns_0_) >>= fun op => returnm (true : bool)
- else if ((eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B1;B1;B0;B1;B1;B1;B1] : mword (6 - 0 + 1)))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B1;B1;B0;B0;B1;B1;B1] : mword (6 - 0 + 1))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ (if sumbool_of_bool w__1 then
+ let _mappingpatterns_0_ : mword 7 := subrange_vec_dec v__463 6 0 in
+ (encdec_uop_backwards _mappingpatterns_0_) >>= fun op => returnm true
+ else if eq_vec (subrange_vec_dec v__463 6 0) ('b"1101111" : mword (6 - 0 + 1)) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"1100111" : mword (6 - 0 + 1))) then
+ returnm true
else
(and_boolM
- (let _mappingpatterns_1_ : mword 3 := subrange_vec_dec v__225 14 12 in
+ (let _mappingpatterns_1_ : mword 3 := subrange_vec_dec v__463 14 12 in
(and_boolM (returnm ((encdec_bop_backwards_matches _mappingpatterns_1_) : bool))
- ((if ((encdec_bop_backwards_matches _mappingpatterns_1_)) then
- (encdec_bop_backwards _mappingpatterns_1_) >>= fun op =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if encdec_bop_backwards_matches _mappingpatterns_1_ then
+ (encdec_bop_backwards _mappingpatterns_1_) >>= fun op => returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
- (returnm ((eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B1;B1;B0;B0;B0;B1;B1] : mword (6 - 0 + 1)))
+ (returnm ((eq_vec (subrange_vec_dec v__463 6 0) ('b"1100011" : mword (6 - 0 + 1)))
: bool))) >>= fun w__4 : bool =>
- (if sumbool_of_bool (w__4) then
- let _mappingpatterns_1_ : mword 3 := subrange_vec_dec v__225 14 12 in
- (encdec_bop_backwards _mappingpatterns_1_) >>= fun op => returnm (true : bool)
+ (if sumbool_of_bool w__4 then
+ let _mappingpatterns_1_ : mword 3 := subrange_vec_dec v__463 14 12 in
+ (encdec_bop_backwards _mappingpatterns_1_) >>= fun op => returnm true
else
(and_boolM
- (let _mappingpatterns_2_ : mword 3 := subrange_vec_dec v__225 14 12 in
+ (let _mappingpatterns_2_ : mword 3 := subrange_vec_dec v__463 14 12 in
(and_boolM (returnm ((encdec_iop_backwards_matches _mappingpatterns_2_) : bool))
- ((if ((encdec_iop_backwards_matches _mappingpatterns_2_)) then
- (encdec_iop_backwards _mappingpatterns_2_) >>= fun op =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if encdec_iop_backwards_matches _mappingpatterns_2_ then
+ (encdec_iop_backwards _mappingpatterns_2_) >>= fun op => returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
- (returnm ((eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword (6 - 0 + 1)))
+ (returnm ((eq_vec (subrange_vec_dec v__463 6 0) ('b"0010011" : mword (6 - 0 + 1)))
: bool))) >>= fun w__7 : bool =>
- (if sumbool_of_bool (w__7) then
- let _mappingpatterns_2_ : mword 3 := subrange_vec_dec v__225 14 12 in
- (encdec_iop_backwards _mappingpatterns_2_) >>= fun op => returnm (true : bool)
+ (if sumbool_of_bool w__7 then
+ let _mappingpatterns_2_ : mword 3 := subrange_vec_dec v__463 14 12 in
+ (encdec_iop_backwards _mappingpatterns_2_) >>= fun op => returnm true
+ else if sumbool_of_bool
+ (andb
+ (let shamt : mword 6 := subrange_vec_dec v__463 25 20 in
+ orb (Z.eqb 64 64) (eq_bit (access_vec_dec shamt 5) B0))
+ (andb (eq_vec (subrange_vec_dec v__463 31 26) ('b"000000" : mword (31 - 26 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0010011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb
+ (let shamt : mword 6 := subrange_vec_dec v__463 25 20 in
+ orb (Z.eqb 64 64) (eq_bit (access_vec_dec shamt 5) B0))
+ (andb (eq_vec (subrange_vec_dec v__463 31 26) ('b"000000" : mword (31 - 26 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0010011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb
+ (let shamt : mword 6 := subrange_vec_dec v__463 25 20 in
+ orb (Z.eqb 64 64) (eq_bit (access_vec_dec shamt 5) B0))
+ (andb (eq_vec (subrange_vec_dec v__463 31 26) ('b"010000" : mword (31 - 26 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0010011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"010" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"011" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"111" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"110" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"100" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0110011" : mword (6 - 0 + 1))))
+ then
+ returnm true
else
- (and_boolMP
- ((let shamt : mword 6 := subrange_vec_dec v__225 25 20 in
- (or_boolMP
- ((returnm (build_ex
- (projT1
- (build_ex
- (Z.eqb 64 64)
- : {_bool : bool & ArithFact (iff (_bool = true) (64 = 64))})))) : M ({_bool : bool & ArithFact (iff (_bool =
- true) (64 = 64))}))
- (build_trivial_ex
- ((bit_to_bool (access_vec_dec shamt 5)) >>= fun w__8 : bool =>
- returnm ((Bool.eqb w__8 false)
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 =
- 64 \/
- simp_0 = true))}))
- : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 = 64 \/
- simp_0 = true))})) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (64 = 64 \/ simp_0 = true))}))
- (build_trivial_ex
- (returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 26)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (31 - 26 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword (6 - 0 + 1)))))
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool = true) (simp_0 =
- true /\
- (64 = 64 \/ simp_1 = true)))})) >>= fun '(existT _ w__10 _ : {_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))}) =>
- (if sumbool_of_bool (w__10) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ (and_boolM
+ (let _mappingpatterns_4_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_3_ : mword 1 := subrange_vec_dec v__463 14 14 in
+ (and_boolM (returnm ((size_bits_backwards_matches _mappingpatterns_4_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_4_ then
+ (size_bits_backwards _mappingpatterns_4_) >>= fun size =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_3_) : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_3_ then
+ (bool_bits_backwards _mappingpatterns_3_) >>= fun is_unsigned =>
+ returnm ((orb (Z.ltb (projT1 (word_width_bytes size)) 8)
+ (andb (negb is_unsigned)
+ ((Z.leb (projT1 (word_width_bytes size)) 8)
+ : bool)))
+ : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((eq_vec (subrange_vec_dec v__463 6 0) ('b"0000011" : mword (6 - 0 + 1)))
+ : bool))) >>= fun w__12 : bool =>
+ (if sumbool_of_bool w__12 then
+ let _mappingpatterns_4_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_3_ : mword 1 := subrange_vec_dec v__463 14 14 in
+ (size_bits_backwards _mappingpatterns_4_) >>= fun size =>
+ (bool_bits_backwards _mappingpatterns_3_) >>= fun is_unsigned => returnm true
else
- (and_boolMP
- ((let shamt : mword 6 := subrange_vec_dec v__225 25 20 in
- (or_boolMP
- ((returnm (build_ex
- (projT1
- (build_ex
- (Z.eqb 64 64)
- : {_bool : bool & ArithFact (iff (_bool = true) (64 = 64))})))) : M ({_bool : bool & ArithFact (iff (_bool =
- true) (64 = 64))}))
- (build_trivial_ex
- ((bit_to_bool (access_vec_dec shamt 5)) >>= fun w__11 : bool =>
- returnm ((Bool.eqb w__11 false)
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 =
- 64 \/
- simp_0 = true))}))
- : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 = 64 \/
- simp_0 = true))})) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (64 = 64 \/ simp_0 = true))}))
- (build_trivial_ex
+ (and_boolM
+ (let _mappingpatterns_5_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ (and_boolM (returnm ((size_bits_backwards_matches _mappingpatterns_5_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_5_ then
+ (size_bits_backwards _mappingpatterns_5_) >>= fun size =>
+ returnm ((Z.leb (projT1 (word_width_bytes size)) 8) : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 26)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (31 - 26 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword (6 - 0 + 1)))))
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))})) >>= fun '(existT _ w__13 _ : {_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))}) =>
- (if sumbool_of_bool (w__13) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ (eq_vec (subrange_vec_dec v__463 14 14)
+ ('b"0"
+ : mword (14 - 14 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0100011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__15 : bool =>
+ (if sumbool_of_bool w__15 then
+ let _mappingpatterns_5_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ (size_bits_backwards _mappingpatterns_5_) >>= fun size => returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__463 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0011011" : mword (6 - 0 + 1)))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"000" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0111011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"001" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0000000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25) ('b"0100000" : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12) ('b"101" : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0) ('b"0011011" : mword (6 - 0 + 1))))))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 28) (Ox"0" : mword (31 - 28 + 1)))
+ (eq_vec (subrange_vec_dec v__463 19 0) (Ox"0000F" : mword (19 - 0 + 1)))
+ then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__463 31 28) (Ox"8" : mword (31 - 28 + 1)))
+ (eq_vec (subrange_vec_dec v__463 19 0) (Ox"0000F" : mword (19 - 0 + 1)))
+ then
+ returnm true
+ else if eq_vec v__463 (Ox"0000100F" : mword 32) then returnm true
+ else if eq_vec v__463 (Ox"00000073" : mword 32) then returnm true
+ else if eq_vec v__463 (Ox"30200073" : mword 32) then returnm true
+ else if eq_vec v__463 (Ox"10200073" : mword 32) then returnm true
+ else if eq_vec v__463 (Ox"00100073" : mword 32) then returnm true
+ else if eq_vec v__463 (Ox"10500073" : mword 32) then returnm true
+ else if andb
+ (eq_vec (subrange_vec_dec v__463 31 25)
+ ('b"0001001"
+ : mword (31 - 25 + 1)))
+ (eq_vec (subrange_vec_dec v__463 14 0)
+ ('b"000000001110011"
+ : mword (14 - 0 + 1))) then
+ returnm true
else
- (and_boolMP
- ((let shamt : mword 6 := subrange_vec_dec v__225 25 20 in
- (or_boolMP
- ((returnm (build_ex
- (projT1
- (build_ex
- (Z.eqb 64 64)
- : {_bool : bool & ArithFact (iff (_bool = true) (64 = 64))})))) : M ({_bool : bool & ArithFact (iff (_bool =
- true) (64 = 64))}))
- (build_trivial_ex
- ((bit_to_bool (access_vec_dec shamt 5)) >>= fun w__14 : bool =>
- returnm ((Bool.eqb w__14 false)
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (64 = 64 \/ simp_0 = true))}))
- : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool = true) (64 = 64 \/
- simp_0 = true))})) : M ({_bool : bool & ArithFact (exists simp_0 , iff (_bool =
- true) (64 = 64 \/ simp_0 = true))}))
- (build_trivial_ex
+ (and_boolM
+ (let _mappingpatterns_8_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_7_ : mword 1 := subrange_vec_dec v__463 25 25 in
+ let _mappingpatterns_6_ : mword 1 := subrange_vec_dec v__463 26 26 in
+ (and_boolM
+ (returnm ((size_bits_backwards_matches _mappingpatterns_8_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_8_ then
+ (size_bits_backwards _mappingpatterns_8_) >>= fun size =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_7_) : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_7_ then
+ (bool_bits_backwards _mappingpatterns_7_) >>= fun rl =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_6_)
+ : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_6_ then
+ (bool_bits_backwards _mappingpatterns_6_) >>= fun aq =>
+ returnm ((Z.leb (projT1 (word_width_bytes size)) 8)
+ : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 26)
- (vec_of_bits [B0;B1;B0;B0;B0;B0] : mword (31 - 26 + 1)))
+ (eq_vec (subrange_vec_dec v__463 31 27)
+ ('b"00010"
+ : mword (31 - 27 + 1)))
(andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B0;B0;B1;B1] : mword (6 - 0 + 1)))))
- : bool))) : M ({_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))})) >>= fun '(existT _ w__16 _ : {_bool : bool & ArithFact (exists simp_0 simp_1 , iff (_bool =
- true) (simp_0 = true /\ (64 = 64 \/ simp_1 = true)))}) =>
- (if sumbool_of_bool (w__16) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B1;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B1;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B1;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B1;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B0] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0] : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1] : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1] : mword (6 - 0 + 1))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ (eq_vec (subrange_vec_dec v__463 24 20)
+ ('b"00000"
+ : mword (24 - 20 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 14)
+ ('b"0"
+ : mword (14 - 14 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0101111"
+ : mword (6 - 0 + 1))))))
+ : bool))) >>= fun w__22 : bool =>
+ (if sumbool_of_bool w__22 then
+ let _mappingpatterns_8_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_7_ : mword 1 := subrange_vec_dec v__463 25 25 in
+ let _mappingpatterns_6_ : mword 1 := subrange_vec_dec v__463 26 26 in
+ (size_bits_backwards _mappingpatterns_8_) >>= fun size =>
+ (bool_bits_backwards _mappingpatterns_7_) >>= fun rl =>
+ (bool_bits_backwards _mappingpatterns_6_) >>= fun aq => returnm true
else
(and_boolM
- (let _mappingpatterns_4_ : mword 2 := subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_3_ : mword 1 := subrange_vec_dec v__225 14 14 in
+ (let _mappingpatterns_9_ : mword 1 := subrange_vec_dec v__463 26 26 in
+ let _mappingpatterns_11_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_10_ : mword 1 := subrange_vec_dec v__463 25 25 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_4_)
- : bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_4_)) then
- (size_bits_backwards _mappingpatterns_4_) >>= fun size =>
+ (returnm ((size_bits_backwards_matches _mappingpatterns_11_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_11_ then
+ (size_bits_backwards _mappingpatterns_11_) >>= fun size =>
(and_boolM
- (returnm ((bool_bits_backwards_matches _mappingpatterns_3_)
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_10_)
: bool))
- ((if ((bool_bits_backwards_matches _mappingpatterns_3_)) then
- (bool_bits_backwards _mappingpatterns_3_) >>= fun is_unsigned =>
- returnm (((orb (Z.ltb (projT1 (word_width_bytes size)) 8)
- (andb (negb is_unsigned)
- ((Z.leb (projT1 (word_width_bytes size)) 8)
- : bool)))
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__18 : bool =>
- returnm ((w__18
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if bool_bits_backwards_matches _mappingpatterns_10_ then
+ (bool_bits_backwards _mappingpatterns_10_) >>= fun rl =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_9_)
+ : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_9_ then
+ (bool_bits_backwards _mappingpatterns_9_) >>= fun aq =>
+ returnm ((Z.leb (projT1 (word_width_bytes size)) 8)
+ : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
: M (bool)))
: M (bool))
- (returnm ((eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B1;B1] : mword (6 - 0 + 1)))
- : bool))) >>= fun w__21 : bool =>
- (if sumbool_of_bool (w__21) then
- let _mappingpatterns_4_ : mword 2 := subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_3_ : mword 1 := subrange_vec_dec v__225 14 14 in
- (size_bits_backwards _mappingpatterns_4_) >>= fun size =>
- (bool_bits_backwards _mappingpatterns_3_) >>= fun is_unsigned =>
- returnm (true
- : bool)
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__463 31 27)
+ ('b"00011"
+ : mword (31 - 27 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 14)
+ ('b"0"
+ : mword (14 - 14 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0101111"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__29 : bool =>
+ (if sumbool_of_bool w__29 then
+ let _mappingpatterns_9_ : mword 1 := subrange_vec_dec v__463 26 26 in
+ let _mappingpatterns_11_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_10_ : mword 1 := subrange_vec_dec v__463 25 25 in
+ (size_bits_backwards _mappingpatterns_11_) >>= fun size =>
+ (bool_bits_backwards _mappingpatterns_10_) >>= fun rl =>
+ (bool_bits_backwards _mappingpatterns_9_) >>= fun aq => returnm true
else
(and_boolM
- (let _mappingpatterns_5_ : mword 2 := subrange_vec_dec v__225 13 12 in
+ (let _mappingpatterns_12_ : mword 5 := subrange_vec_dec v__463 31 27 in
+ let _mappingpatterns_15_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_14_ : mword 1 := subrange_vec_dec v__463 25 25 in
+ let _mappingpatterns_13_ : mword 1 := subrange_vec_dec v__463 26 26 in
+ let _mappingpatterns_12_ : mword 5 := subrange_vec_dec v__463 31 27 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_5_)
- : bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_5_)) then
- (size_bits_backwards _mappingpatterns_5_) >>= fun size =>
- returnm (((Z.leb (projT1 (word_width_bytes size)) 8)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ (returnm ((size_bits_backwards_matches _mappingpatterns_15_) : bool))
+ ((if size_bits_backwards_matches _mappingpatterns_15_ then
+ (size_bits_backwards _mappingpatterns_15_) >>= fun size =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches _mappingpatterns_14_)
+ : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_14_ then
+ (bool_bits_backwards _mappingpatterns_14_) >>= fun rl =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches
+ _mappingpatterns_13_)
+ : bool))
+ ((if bool_bits_backwards_matches _mappingpatterns_13_
+ then
+ (bool_bits_backwards _mappingpatterns_13_) >>= fun aq =>
+ (and_boolM
+ (returnm ((encdec_amoop_backwards_matches
+ _mappingpatterns_12_)
+ : bool))
+ ((if encdec_amoop_backwards_matches
+ _mappingpatterns_12_ then
+ (encdec_amoop_backwards _mappingpatterns_12_) >>= fun op =>
+ returnm ((Z.leb
+ (projT1
+ (word_width_bytes size)) 8)
+ : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 14 14)
- (vec_of_bits [B0] : mword (14 - 14 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B0;B0;B0;B1;B1] : mword (6 - 0 + 1))))
- : bool))) >>= fun w__24 : bool =>
- (if sumbool_of_bool (w__24) then
- let _mappingpatterns_5_ : mword 2 := subrange_vec_dec v__225 13 12 in
- (size_bits_backwards _mappingpatterns_5_) >>= fun size =>
- returnm (true
- : bool)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B0]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1)))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B0]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B0]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B0;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B1;B0;B0;B0;B0;B0]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 12)
- (vec_of_bits [B1;B0;B1]
- : mword (14 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B0;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1))))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 28)
- (vec_of_bits [B0;B0;B0;B0] : mword (31 - 28 + 1)))
- (eq_vec (subrange_vec_dec v__225 19 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B1]
- : mword (19 - 0 + 1))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 28)
- (vec_of_bits [B1;B0;B0;B0] : mword (31 - 28 + 1)))
- (eq_vec (subrange_vec_dec v__225 19 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B1]
- : mword (19 - 0 + 1))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((eq_vec v__225
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;
- B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B1]
- : mword 32))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((eq_vec v__225
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((eq_vec v__225
- (vec_of_bits [B0;B0;B1;B1;B0;B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((eq_vec v__225
- (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((eq_vec v__225
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((eq_vec v__225
- (vec_of_bits [B0;B0;B0;B1;B0;B0;B0;B0;B0;B1;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B1;B0;B0;B1] : mword (31 - 25 + 1)))
- (eq_vec (subrange_vec_dec v__225 14 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword (14 - 0 + 1))))) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ (eq_vec (subrange_vec_dec v__463 14 14)
+ ('b"0"
+ : mword (14 - 14 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0101111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__38 : bool =>
+ (if sumbool_of_bool w__38 then
+ let _mappingpatterns_12_ : mword 5 := subrange_vec_dec v__463 31 27 in
+ let _mappingpatterns_15_ : mword 2 := subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_14_ : mword 1 := subrange_vec_dec v__463 25 25 in
+ let _mappingpatterns_13_ : mword 1 := subrange_vec_dec v__463 26 26 in
+ let _mappingpatterns_12_ : mword 5 := subrange_vec_dec v__463 31 27 in
+ (size_bits_backwards _mappingpatterns_15_) >>= fun size =>
+ (bool_bits_backwards _mappingpatterns_14_) >>= fun rl =>
+ (bool_bits_backwards _mappingpatterns_13_) >>= fun aq =>
+ (encdec_amoop_backwards _mappingpatterns_12_) >>= fun op => returnm true
else
(and_boolM
- (let _mappingpatterns_8_ : mword 2 := subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_7_ : mword 1 := subrange_vec_dec v__225 25 25 in
- let _mappingpatterns_6_ : mword 1 := subrange_vec_dec v__225 26 26 in
+ (let _mappingpatterns_16_ : mword 3 := subrange_vec_dec v__463 14 12 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_8_)
+ (returnm ((encdec_mul_op_backwards_matches _mappingpatterns_16_)
: bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_8_)) then
- (size_bits_backwards _mappingpatterns_8_) >>= fun size =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches _mappingpatterns_7_)
- : bool))
- ((if ((bool_bits_backwards_matches _mappingpatterns_7_))
- then
- (bool_bits_backwards _mappingpatterns_7_) >>= fun rl =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_6_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_6_)) then
- (bool_bits_backwards _mappingpatterns_6_) >>= fun aq =>
- returnm (((Z.leb (projT1 (word_width_bytes size))
- 8)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__26 : bool =>
- returnm ((w__26
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__28 : bool =>
- returnm ((w__28
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if encdec_mul_op_backwards_matches _mappingpatterns_16_ then
+ (encdec_mul_op_backwards _mappingpatterns_16_) >>= fun '(high, signed1, signed2) =>
+ returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 27)
- (vec_of_bits [B0;B0;B0;B1;B0] : mword (31 - 27 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 24 20)
- (vec_of_bits [B0;B0;B0;B0;B0]
- : mword (24 - 20 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 14)
- (vec_of_bits [B0] : mword (14 - 14 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1]
- : mword (6 - 0 + 1))))))
- : bool))) >>= fun w__31 : bool =>
- (if sumbool_of_bool (w__31) then
- let _mappingpatterns_8_ : mword 2 := subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_7_ : mword 1 := subrange_vec_dec v__225 25 25 in
- let _mappingpatterns_6_ : mword 1 := subrange_vec_dec v__225 26 26 in
- (size_bits_backwards _mappingpatterns_8_) >>= fun size =>
- (bool_bits_backwards _mappingpatterns_7_) >>= fun rl =>
- (bool_bits_backwards _mappingpatterns_6_) >>= fun aq =>
- returnm (true
- : bool)
+ (eq_vec (subrange_vec_dec v__463 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0110011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__41 : bool =>
+ (if sumbool_of_bool w__41 then
+ let _mappingpatterns_16_ : mword 3 := subrange_vec_dec v__463 14 12 in
+ (encdec_mul_op_backwards _mappingpatterns_16_) >>= fun '(high, signed1, signed2) =>
+ returnm true
else
(and_boolM
- (let _mappingpatterns_9_ : mword 1 := subrange_vec_dec v__225 26 26 in
- let _mappingpatterns_11_ : mword 2 := subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_10_ : mword 1 := subrange_vec_dec v__225 25 25 in
+ (let _mappingpatterns_17_ : mword 1 :=
+ subrange_vec_dec v__463 12 12 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_11_)
+ (returnm ((bool_not_bits_backwards_matches _mappingpatterns_17_)
: bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_11_)) then
- (size_bits_backwards _mappingpatterns_11_) >>= fun size =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_10_)
- : bool))
- ((if ((bool_bits_backwards_matches _mappingpatterns_10_))
- then
- (bool_bits_backwards _mappingpatterns_10_) >>= fun rl =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_9_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_9_)) then
- (bool_bits_backwards _mappingpatterns_9_) >>= fun aq =>
- returnm (((Z.leb
- (projT1
- (word_width_bytes size)) 8)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool))) >>= fun w__33 : bool =>
- returnm ((w__33
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__35 : bool =>
- returnm ((w__35
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if bool_not_bits_backwards_matches _mappingpatterns_17_ then
+ (bool_not_bits_backwards _mappingpatterns_17_) >>= fun s =>
+ returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 27)
- (vec_of_bits [B0;B0;B0;B1;B1]
- : mword (31 - 27 + 1)))
+ (eq_vec (subrange_vec_dec v__463 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
(andb
- (eq_vec (subrange_vec_dec v__225 14 14)
- (vec_of_bits [B0] : mword (14 - 14 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1]
- : mword (6 - 0 + 1)))))
- : bool))) >>= fun w__38 : bool =>
- (if sumbool_of_bool (w__38) then
- let _mappingpatterns_9_ : mword 1 := subrange_vec_dec v__225 26 26 in
- let _mappingpatterns_11_ : mword 2 := subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_10_ : mword 1 := subrange_vec_dec v__225 25 25 in
- (size_bits_backwards _mappingpatterns_11_) >>= fun size =>
- (bool_bits_backwards _mappingpatterns_10_) >>= fun rl =>
- (bool_bits_backwards _mappingpatterns_9_) >>= fun aq =>
- returnm (true
- : bool)
+ (eq_vec (subrange_vec_dec v__463 14 13)
+ ('b"10"
+ : mword (14 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0110011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__44 : bool =>
+ (if sumbool_of_bool w__44 then
+ let _mappingpatterns_17_ : mword 1 := subrange_vec_dec v__463 12 12 in
+ (bool_not_bits_backwards _mappingpatterns_17_) >>= fun s =>
+ returnm true
else
(and_boolM
- (let _mappingpatterns_12_ : mword 5 :=
- subrange_vec_dec v__225 31 27 in
- let _mappingpatterns_15_ : mword 2 :=
- subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_14_ : mword 1 :=
- subrange_vec_dec v__225 25 25 in
- let _mappingpatterns_13_ : mword 1 :=
- subrange_vec_dec v__225 26 26 in
- let _mappingpatterns_12_ : mword 5 :=
- subrange_vec_dec v__225 31 27 in
+ (let _mappingpatterns_18_ : mword 1 :=
+ subrange_vec_dec v__463 12 12 in
(and_boolM
- (returnm ((size_bits_backwards_matches _mappingpatterns_15_)
+ (returnm ((bool_not_bits_backwards_matches
+ _mappingpatterns_18_)
: bool))
- ((if ((size_bits_backwards_matches _mappingpatterns_15_))
+ ((if bool_not_bits_backwards_matches _mappingpatterns_18_
then
- (size_bits_backwards _mappingpatterns_15_) >>= fun size =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_14_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_14_)) then
- (bool_bits_backwards _mappingpatterns_14_) >>= fun rl =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_13_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_13_)) then
- (bool_bits_backwards _mappingpatterns_13_) >>= fun aq =>
- (and_boolM
- (returnm ((encdec_amoop_backwards_matches
- _mappingpatterns_12_)
- : bool))
- ((if ((encdec_amoop_backwards_matches
- _mappingpatterns_12_)) then
- (encdec_amoop_backwards
- _mappingpatterns_12_) >>= fun op =>
- returnm (((Z.leb
- (projT1
- (word_width_bytes size))
- 8)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool))) >>= fun w__40 : bool =>
- returnm ((w__40
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool))) >>= fun w__42 : bool =>
- returnm ((w__42
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
- : M (bool))) >>= fun w__44 : bool =>
- returnm ((w__44
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ (bool_not_bits_backwards _mappingpatterns_18_) >>= fun s =>
+ returnm (true : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 14 14)
- (vec_of_bits [B0] : mword (14 - 14 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B0;B1;B1;B1;B1]
- : mword (6 - 0 + 1))))
+ (eq_vec (subrange_vec_dec v__463 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 13)
+ ('b"11"
+ : mword (14 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0110011"
+ : mword (6 - 0 + 1)))))
: bool))) >>= fun w__47 : bool =>
- (if sumbool_of_bool (w__47) then
- let _mappingpatterns_12_ : mword 5 :=
- subrange_vec_dec v__225 31 27 in
- let _mappingpatterns_15_ : mword 2 :=
- subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_14_ : mword 1 :=
- subrange_vec_dec v__225 25 25 in
- let _mappingpatterns_13_ : mword 1 :=
- subrange_vec_dec v__225 26 26 in
- let _mappingpatterns_12_ : mword 5 :=
- subrange_vec_dec v__225 31 27 in
- (size_bits_backwards _mappingpatterns_15_) >>= fun size =>
- (bool_bits_backwards _mappingpatterns_14_) >>= fun rl =>
- (bool_bits_backwards _mappingpatterns_13_) >>= fun aq =>
- (encdec_amoop_backwards _mappingpatterns_12_) >>= fun op =>
- returnm (true
- : bool)
+ (if sumbool_of_bool w__47 then
+ let _mappingpatterns_18_ : mword 1 :=
+ subrange_vec_dec v__463 12 12 in
+ (bool_not_bits_backwards _mappingpatterns_18_) >>= fun s =>
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb
+ (eq_vec (subrange_vec_dec v__463 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0111011"
+ : mword (6 - 0 + 1)))))) then
+ returnm true
else
(and_boolM
- (let _mappingpatterns_16_ : mword 3 :=
- subrange_vec_dec v__225 14 12 in
+ (let _mappingpatterns_19_ : mword 1 :=
+ subrange_vec_dec v__463 12 12 in
(and_boolM
- (returnm ((encdec_mul_op_backwards_matches
- _mappingpatterns_16_)
+ (returnm ((bool_not_bits_backwards_matches
+ _mappingpatterns_19_)
: bool))
- ((if ((encdec_mul_op_backwards_matches
- _mappingpatterns_16_)) then
- (encdec_mul_op_backwards _mappingpatterns_16_) >>= fun '(high, signed1, signed2) =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if bool_not_bits_backwards_matches _mappingpatterns_19_
+ then
+ (bool_not_bits_backwards _mappingpatterns_19_) >>= fun s =>
+ returnm ((Z.eqb 64 64) : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1]
- : mword (6 - 0 + 1))))
+ (eq_vec (subrange_vec_dec v__463 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec (subrange_vec_dec v__463 14 13)
+ ('b"10"
+ : mword (14 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0111011"
+ : mword (6 - 0 + 1)))))
: bool))) >>= fun w__50 : bool =>
- (if sumbool_of_bool (w__50) then
- let _mappingpatterns_16_ : mword 3 :=
- subrange_vec_dec v__225 14 12 in
- (encdec_mul_op_backwards _mappingpatterns_16_) >>= fun '(high, signed1, signed2) =>
- returnm (true
- : bool)
+ (if sumbool_of_bool w__50 then
+ let _mappingpatterns_19_ : mword 1 :=
+ subrange_vec_dec v__463 12 12 in
+ (bool_not_bits_backwards _mappingpatterns_19_) >>= fun s =>
+ returnm true
else
(and_boolM
- (let _mappingpatterns_17_ : mword 1 :=
- subrange_vec_dec v__225 12 12 in
+ (let _mappingpatterns_20_ : mword 1 :=
+ subrange_vec_dec v__463 12 12 in
(and_boolM
(returnm ((bool_not_bits_backwards_matches
- _mappingpatterns_17_)
+ _mappingpatterns_20_)
: bool))
- ((if ((bool_not_bits_backwards_matches
- _mappingpatterns_17_)) then
- (bool_not_bits_backwards _mappingpatterns_17_) >>= fun s =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if bool_not_bits_backwards_matches
+ _mappingpatterns_20_ then
+ (bool_not_bits_backwards _mappingpatterns_20_) >>= fun s =>
+ returnm ((Z.eqb 64 64) : bool)
+ else returnm false)
: M (bool)))
: M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
+ (eq_vec (subrange_vec_dec v__463 31 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
(andb
- (eq_vec (subrange_vec_dec v__225 14 13)
- (vec_of_bits [B1;B0]
- : mword (14 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1]
- : mword (6 - 0 + 1)))))
+ (eq_vec (subrange_vec_dec v__463 14 13)
+ ('b"11"
+ : mword (14 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0111011"
+ : mword (6 - 0 + 1)))))
: bool))) >>= fun w__53 : bool =>
- (if sumbool_of_bool (w__53) then
- let _mappingpatterns_17_ : mword 1 :=
- subrange_vec_dec v__225 12 12 in
- (bool_not_bits_backwards _mappingpatterns_17_) >>= fun s =>
- returnm (true
- : bool)
+ (if sumbool_of_bool w__53 then
+ let _mappingpatterns_20_ : mword 1 :=
+ subrange_vec_dec v__463 12 12 in
+ (bool_not_bits_backwards _mappingpatterns_20_) >>= fun s =>
+ returnm true
else
(and_boolM
- (let _mappingpatterns_18_ : mword 1 :=
- subrange_vec_dec v__225 12 12 in
+ (let _mappingpatterns_22_ : mword 2 :=
+ subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_21_ : mword 1 :=
+ subrange_vec_dec v__463 14 14 in
(and_boolM
- (returnm ((bool_not_bits_backwards_matches
- _mappingpatterns_18_)
+ (returnm ((encdec_csrop_backwards_matches
+ _mappingpatterns_22_)
: bool))
- ((if ((bool_not_bits_backwards_matches
- _mappingpatterns_18_)) then
- (bool_not_bits_backwards _mappingpatterns_18_) >>= fun s =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool = true))})))
+ ((if encdec_csrop_backwards_matches
+ _mappingpatterns_22_ then
+ (encdec_csrop_backwards _mappingpatterns_22_) >>= fun op =>
+ (and_boolM
+ (returnm ((bool_bits_backwards_matches
+ _mappingpatterns_21_)
+ : bool))
+ ((if bool_bits_backwards_matches
+ _mappingpatterns_21_ then
+ (bool_bits_backwards _mappingpatterns_21_) >>= fun is_imm =>
+ returnm (true : bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool)
+ else returnm false)
: M (bool)))
: M (bool))
- (returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__225 14 13)
- (vec_of_bits [B1;B1]
- : mword (14 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B0;B0;B1;B1]
- : mword (6 - 0 + 1)))))
- : bool))) >>= fun w__56 : bool =>
- (if sumbool_of_bool (w__56) then
- let _mappingpatterns_18_ : mword 1 :=
- subrange_vec_dec v__225 12 12 in
- (bool_not_bits_backwards _mappingpatterns_18_) >>= fun s =>
- returnm (true
- : bool)
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec
- (subrange_vec_dec
- v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec
- (subrange_vec_dec
- v__225 14 12)
- (vec_of_bits [B0;B0;B0]
- : mword (14 - 12 + 1)))
- (eq_vec
- (subrange_vec_dec
- v__225 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1)))))))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ (returnm ((eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"1110011"
+ : mword (6 - 0 + 1)))
+ : bool))) >>= fun w__58 : bool =>
+ (if sumbool_of_bool w__58 then
+ let _mappingpatterns_22_ : mword 2 :=
+ subrange_vec_dec v__463 13 12 in
+ let _mappingpatterns_21_ : mword 1 :=
+ subrange_vec_dec v__463 14 14 in
+ (encdec_csrop_backwards _mappingpatterns_22_) >>= fun op =>
+ (bool_bits_backwards _mappingpatterns_21_) >>= fun is_imm =>
+ returnm true
+ else if eq_vec v__463 (Ox"00200073" : mword 32) then
+ returnm true
else
- (and_boolM
- (let _mappingpatterns_19_ : mword 1 :=
- subrange_vec_dec v__225 12 12 in
- (and_boolM
- (returnm ((bool_not_bits_backwards_matches
- _mappingpatterns_19_)
- : bool))
- ((if ((bool_not_bits_backwards_matches
- _mappingpatterns_19_)) then
- (bool_not_bits_backwards _mappingpatterns_19_) >>= fun s =>
- returnm (((Z.eqb 64 64)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool)))
- : M (bool))
+ (and_boolM ((is_RV32F_or_RV64F tt) : M (bool))
(returnm ((andb
- (eq_vec (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec
- (subrange_vec_dec v__225 14 13)
- (vec_of_bits [B1;B0]
- : mword (14 - 13 + 1)))
- (eq_vec
- (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1)))))
- : bool))) >>= fun w__59 : bool =>
- (if sumbool_of_bool (w__59) then
- let _mappingpatterns_19_ : mword 1 :=
- subrange_vec_dec v__225 12 12 in
- (bool_not_bits_backwards _mappingpatterns_19_) >>= fun s =>
- returnm (true
- : bool)
+ (eq_vec (subrange_vec_dec v__463 14 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__463 6 0)
+ ('b"0000111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__60 : bool =>
+ (if sumbool_of_bool w__60 then returnm true
else
- (and_boolM
- (let _mappingpatterns_20_ : mword 1 :=
- subrange_vec_dec v__225 12 12 in
- (and_boolM
- (returnm ((bool_not_bits_backwards_matches
- _mappingpatterns_20_)
- : bool))
- ((if ((bool_not_bits_backwards_matches
- _mappingpatterns_20_)) then
- (bool_not_bits_backwards
- _mappingpatterns_20_) >>= fun s =>
- returnm (((Z.eqb 64 64)
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool)))
- : M (bool))
+ (and_boolM ((is_RV32D_or_RV64D tt) : M (bool))
(returnm ((andb
(eq_vec
- (subrange_vec_dec v__225 31 25)
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B1]
- : mword (31 - 25 + 1)))
- (andb
- (eq_vec
- (subrange_vec_dec v__225 14 13)
- (vec_of_bits [B1;B1]
- : mword (14 - 13 + 1)))
- (eq_vec
- (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B0;B1;B1;B1;B0;B1;B1]
- : mword (6 - 0 + 1)))))
+ (subrange_vec_dec v__463 14 12)
+ ('b"011"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec v__463 6 0)
+ ('b"0000111"
+ : mword (6 - 0 + 1))))
: bool))) >>= fun w__62 : bool =>
- (if sumbool_of_bool (w__62) then
- let _mappingpatterns_20_ : mword 1 :=
- subrange_vec_dec v__225 12 12 in
- (bool_not_bits_backwards _mappingpatterns_20_) >>= fun s =>
- returnm (true
- : bool)
+ (if sumbool_of_bool w__62 then returnm true
else
- (and_boolM
- (let _mappingpatterns_22_ : mword 2 :=
- subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_21_ : mword 1 :=
- subrange_vec_dec v__225 14 14 in
+ (and_boolM ((is_RV32F_or_RV64F tt) : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec v__463 14 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec v__463 6 0)
+ ('b"0100111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__64 : bool =>
+ (if sumbool_of_bool w__64 then returnm true
+ else
(and_boolM
- (returnm ((encdec_csrop_backwards_matches
- _mappingpatterns_22_)
- : bool))
- ((if ((encdec_csrop_backwards_matches
- _mappingpatterns_22_)) then
- (encdec_csrop_backwards
- _mappingpatterns_22_) >>= fun op =>
- (and_boolM
- (returnm ((bool_bits_backwards_matches
- _mappingpatterns_21_)
- : bool))
- ((if ((bool_bits_backwards_matches
- _mappingpatterns_21_)) then
- (bool_bits_backwards
- _mappingpatterns_21_) >>= fun is_imm =>
- returnm ((true
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool))) >>= fun w__64 : bool =>
- returnm ((w__64
- : bool)
- : bool)
- else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
- : M (bool)))
+ ((is_RV32D_or_RV64D tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec v__463 14
+ 12)
+ ('b"011"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec v__463 6
+ 0)
+ ('b"0100111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__66 : bool =>
+ (if sumbool_of_bool w__66 then returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_23_ : mword 3 :=
+ subrange_vec_dec v__463 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_23_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_23_ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_23_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec v__463
+ 26 25)
+ ('b"00"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec v__463
+ 6 0)
+ ('b"1000011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__70 : bool =>
+ (if sumbool_of_bool w__70 then
+ let _mappingpatterns_23_ : mword 3 :=
+ subrange_vec_dec v__463 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_23_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_24_ : mword 3 :=
+ subrange_vec_dec v__463 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_24_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_24_ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_24_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463 26 25)
+ ('b"00"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463 6 0)
+ ('b"1000111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__74 : bool =>
+ (if sumbool_of_bool w__74 then
+ let _mappingpatterns_24_ : mword 3 :=
+ subrange_vec_dec v__463 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_24_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_25_ : mword 3 :=
+ subrange_vec_dec v__463 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_25_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_25_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_25_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463 26 25)
+ ('b"00"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463 6 0)
+ ('b"1001011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__78 : bool =>
+ (if sumbool_of_bool w__78 then
+ let _mappingpatterns_25_ : mword 3 :=
+ subrange_vec_dec v__463 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_25_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_26_ : mword 3 :=
+ subrange_vec_dec v__463 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_26_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_26_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_26_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463 26 25)
+ ('b"00"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463 6 0)
+ ('b"1001111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__82 : bool =>
+ (if sumbool_of_bool w__82 then
+ let _mappingpatterns_26_ : mword 3 :=
+ subrange_vec_dec v__463 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_26_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_27_ : mword 3 :=
+ subrange_vec_dec v__463 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_27_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_27_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_27_) >>= fun rm =>
+ (is_RV32F_or_RV64F tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463 31
+ 25)
+ ('b"0000000"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__86 : bool =>
+ (if sumbool_of_bool w__86 then
+ let _mappingpatterns_27_ : mword 3 :=
+ subrange_vec_dec v__463 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_27_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_28_ : mword 3 :=
+ subrange_vec_dec v__463
+ 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_28_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_28_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_28_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31 25)
+ ('b"0000100"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__90 : bool =>
+ (if sumbool_of_bool
+ w__90 then
+ let _mappingpatterns_28_ : mword 3 :=
+ subrange_vec_dec v__463
+ 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_28_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_29_ : mword 3 :=
+ subrange_vec_dec
+ v__463 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_29_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_29_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_29_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0001000"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__94 : bool =>
+ (if sumbool_of_bool
+ w__94 then
+ let _mappingpatterns_29_ : mword 3 :=
+ subrange_vec_dec
+ v__463 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_29_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_30_ : mword 3 :=
+ subrange_vec_dec
+ v__463 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_30_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_30_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_30_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0001100"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__98 : bool =>
+ (if sumbool_of_bool
+ w__98 then
+ let _mappingpatterns_30_ : mword 3 :=
+ subrange_vec_dec
+ v__463 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_30_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_31_ : mword 3 :=
+ subrange_vec_dec
+ v__463 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_31_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_31_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_31_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"580"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__102 : bool =>
+ (if sumbool_of_bool
+ w__102 then
+ let _mappingpatterns_31_ : mword 3 :=
+ subrange_vec_dec
+ v__463 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_31_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_32_ : mword 3 :=
+ subrange_vec_dec
+ v__463 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_32_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_32_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_32_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"C00"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__106 : bool =>
+ (if sumbool_of_bool
+ w__106 then
+ let _mappingpatterns_32_ : mword 3 :=
+ subrange_vec_dec
+ v__463 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_32_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_33_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_33_)
+ : bool))
+ ((if encdec_rounding_mode_backwards_matches
+ _mappingpatterns_33_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_33_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"C01"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__110 : bool =>
+ (if sumbool_of_bool
+ w__110
+ then
+ let _mappingpatterns_33_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_33_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_34_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_34_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_34_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_34_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"D00"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__114 : bool =>
+ (if sumbool_of_bool
+ w__114
+ then
+ let _mappingpatterns_34_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_34_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_35_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_35_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_35_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_35_) >>= fun rm =>
+ (is_RV32F_or_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"D01"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__118 : bool =>
+ (if sumbool_of_bool
+ w__118
+ then
+ let _mappingpatterns_35_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_35_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_36_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_36_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_36_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_36_) >>= fun rm =>
+ (is_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"C02"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__122 : bool =>
+ (if sumbool_of_bool
+ w__122
+ then
+ let _mappingpatterns_36_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_36_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_37_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_37_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_37_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_37_) >>= fun rm =>
+ (is_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"C03"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__126 : bool =>
+ (if
+ sumbool_of_bool
+ w__126
+ then
+ let _mappingpatterns_37_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_37_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_38_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_38_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_38_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_38_) >>= fun rm =>
+ (is_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"D02"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__130 : bool =>
+ (if
+ sumbool_of_bool
+ w__130
+ then
+ let _mappingpatterns_38_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_38_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_39_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_39_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_39_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_39_) >>= fun rm =>
+ (is_RV64F
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"D03"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__134 : bool =>
+ (if
+ sumbool_of_bool
+ w__134
+ then
+ let _mappingpatterns_39_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_39_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__136 : bool =>
+ (if
+ sumbool_of_bool
+ w__136
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__138 : bool =>
+ (if
+ sumbool_of_bool
+ w__138
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__140 : bool =>
+ (if
+ sumbool_of_bool
+ w__140
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010100"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__142 : bool =>
+ (if
+ sumbool_of_bool
+ w__142
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010100"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__144 : bool =>
+ (if
+ sumbool_of_bool
+ w__144
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"1010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__146 : bool =>
+ (if
+ sumbool_of_bool
+ w__146
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"1010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__148 : bool =>
+ (if
+ sumbool_of_bool
+ w__148
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32F_or_RV64F
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"1010000"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__150 : bool =>
+ (if
+ sumbool_of_bool
+ w__150
+ then
+ returnm true
+ else
+ (and_boolM
+ ((haveFExt
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"E00"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__152 : bool =>
+ (if
+ sumbool_of_bool
+ w__152
+ then
+ returnm true
+ else
+ (and_boolM
+ ((haveFExt
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"E00"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__154 : bool =>
+ (if
+ sumbool_of_bool
+ w__154
+ then
+ returnm true
+ else
+ (and_boolM
+ ((haveFExt
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"F00"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__156 : bool =>
+ (if
+ sumbool_of_bool
+ w__156
+ then
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_40_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_40_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_40_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_40_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 26
+ 25)
+ ('b"01"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1000011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__160 : bool =>
+ (if
+ sumbool_of_bool
+ w__160
+ then
+ let _mappingpatterns_40_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_40_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_41_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_41_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_41_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_41_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 26
+ 25)
+ ('b"01"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1000111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__164 : bool =>
+ (if
+ sumbool_of_bool
+ w__164
+ then
+ let _mappingpatterns_41_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_41_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_42_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_42_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_42_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_42_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 26
+ 25)
+ ('b"01"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1001011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__168 : bool =>
+ (if
+ sumbool_of_bool
+ w__168
+ then
+ let _mappingpatterns_42_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_42_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_43_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_43_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_43_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_43_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 26
+ 25)
+ ('b"01"
+ : mword (26 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1001111"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__172 : bool =>
+ (if
+ sumbool_of_bool
+ w__172
+ then
+ let _mappingpatterns_43_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_43_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_44_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_44_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_44_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_44_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0000001"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__176 : bool =>
+ (if
+ sumbool_of_bool
+ w__176
+ then
+ let _mappingpatterns_44_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_44_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_45_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_45_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_45_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_45_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0000101"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__180 : bool =>
+ (if
+ sumbool_of_bool
+ w__180
+ then
+ let _mappingpatterns_45_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_45_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_46_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_46_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_46_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_46_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0001001"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__184 : bool =>
+ (if
+ sumbool_of_bool
+ w__184
+ then
+ let _mappingpatterns_46_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_46_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_47_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_47_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_47_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_47_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0001101"
+ : mword (31 - 25 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__188 : bool =>
+ (if
+ sumbool_of_bool
+ w__188
+ then
+ let _mappingpatterns_47_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_47_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_48_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_48_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_48_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_48_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"5A0"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__192 : bool =>
+ (if
+ sumbool_of_bool
+ w__192
+ then
+ let _mappingpatterns_48_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_48_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_49_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_49_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_49_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_49_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"C20"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__196 : bool =>
+ (if
+ sumbool_of_bool
+ w__196
+ then
+ let _mappingpatterns_49_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_49_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_50_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_50_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_50_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_50_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"C21"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__200 : bool =>
+ (if
+ sumbool_of_bool
+ w__200
+ then
+ let _mappingpatterns_50_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_50_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_51_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_51_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_51_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_51_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"D20"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__204 : bool =>
+ (if
+ sumbool_of_bool
+ w__204
+ then
+ let _mappingpatterns_51_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_51_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_52_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_52_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_52_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_52_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"D21"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__208 : bool =>
+ (if
+ sumbool_of_bool
+ w__208
+ then
+ let _mappingpatterns_52_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_52_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_53_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_53_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_53_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_53_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"401"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__212 : bool =>
+ (if
+ sumbool_of_bool
+ w__212
+ then
+ let _mappingpatterns_53_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_53_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_54_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_54_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_54_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_54_) >>= fun rm =>
+ (is_RV32D_or_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"420"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__216 : bool =>
+ (if
+ sumbool_of_bool
+ w__216
+ then
+ let _mappingpatterns_54_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_54_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_55_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_55_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_55_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_55_) >>= fun rm =>
+ (is_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"C22"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__220 : bool =>
+ (if
+ sumbool_of_bool
+ w__220
+ then
+ let _mappingpatterns_55_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_55_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_56_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_56_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_56_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_56_) >>= fun rm =>
+ (is_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"C23"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__224 : bool =>
+ (if
+ sumbool_of_bool
+ w__224
+ then
+ let _mappingpatterns_56_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_56_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_57_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_57_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_57_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_57_) >>= fun rm =>
+ (is_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"D22"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__228 : bool =>
+ (if
+ sumbool_of_bool
+ w__228
+ then
+ let _mappingpatterns_57_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_57_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ (let _mappingpatterns_58_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (and_boolM
+ (returnm ((encdec_rounding_mode_backwards_matches
+ _mappingpatterns_58_)
+ : bool))
+ ((if
+ encdec_rounding_mode_backwards_matches
+ _mappingpatterns_58_
+ then
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_58_) >>= fun rm =>
+ (is_RV64D
+ tt)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool)))
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"D23"
+ : mword (31 - 20 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1))))
+ : bool))) >>= fun w__232 : bool =>
+ (if
+ sumbool_of_bool
+ w__232
+ then
+ let _mappingpatterns_58_ : mword 3 :=
+ subrange_vec_dec
+ v__463
+ 14
+ 12 in
+ (encdec_rounding_mode_backwards
+ _mappingpatterns_58_) >>= fun rm =>
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__234 : bool =>
+ (if
+ sumbool_of_bool
+ w__234
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__236 : bool =>
+ (if
+ sumbool_of_bool
+ w__236
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__238 : bool =>
+ (if
+ sumbool_of_bool
+ w__238
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010101"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__240 : bool =>
+ (if
+ sumbool_of_bool
+ w__240
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"0010101"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__242 : bool =>
+ (if
+ sumbool_of_bool
+ w__242
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"1010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"010"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__244 : bool =>
+ (if
+ sumbool_of_bool
+ w__244
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"1010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__246 : bool =>
+ (if
+ sumbool_of_bool
+ w__246
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV32D_or_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 25)
+ ('b"1010001"
+ : mword (31 - 25 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__248 : bool =>
+ (if
+ sumbool_of_bool
+ w__248
+ then
+ returnm true
+ else
+ (and_boolM
+ ((haveDExt
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"E20"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"001"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__250 : bool =>
+ (if
+ sumbool_of_bool
+ w__250
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"E20"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__252 : bool =>
+ (if
+ sumbool_of_bool
+ w__252
+ then
+ returnm true
+ else
+ (and_boolM
+ ((is_RV64D
+ tt)
+ : M (bool))
+ (returnm ((andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 31
+ 20)
+ (Ox"F20"
+ : mword (31 - 20 + 1)))
+ (andb
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 14
+ 12)
+ ('b"000"
+ : mword (14 - 12 + 1)))
+ (eq_vec
+ (subrange_vec_dec
+ v__463
+ 6
+ 0)
+ ('b"1010011"
+ : mword (6 - 0 + 1)))))
+ : bool))) >>= fun w__254 : bool =>
+ returnm (if
+ sumbool_of_bool
+ w__254
+ then
+ true
+ else
+ true))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
: M (bool))
- (returnm ((eq_vec
- (subrange_vec_dec v__225 6 0)
- (vec_of_bits [B1;B1;B1;B0;B0;B1;B1]
- : mword (6 - 0 + 1)))
- : bool))) >>= fun w__67 : bool =>
- (if sumbool_of_bool (w__67) then
- let _mappingpatterns_22_ : mword 2 :=
- subrange_vec_dec v__225 13 12 in
- let _mappingpatterns_21_ : mword 1 :=
- subrange_vec_dec v__225 14 14 in
- (encdec_csrop_backwards _mappingpatterns_22_) >>= fun op =>
- (bool_bits_backwards _mappingpatterns_21_) >>= fun is_imm =>
- returnm (true
- : bool)
- else
- returnm ((if ((eq_vec v__225
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;
- B0;B0;B0;B0;B1;B1;B1;B0;B0;B1;B1]
- : mword 32))) then
- true
- else true)
- : bool))
: M (bool))
: M (bool))
: M (bool))
@@ -20420,1637 +26833,1976 @@ Definition encdec_backwards_matches (arg_ : mword 32)
: M (bool))
: M (bool).
-Definition encdec_compressed_forwards (arg_ : ast)
-: M (mword 16) :=
-
+Definition encdec_compressed_forwards (arg_ : ast) : M (mword 16) :=
(match arg_ with
- | C_NOP (tt) =>
- returnm ((concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B0;B1] : mword 2)))))
- : mword (3 + (1 + (5 + (5 + 2)))))
- | C_ADDI4SPN ((rd, v__438)) =>
- (if ((let nz96 : bits 4 := subrange_vec_dec v__438 7 4 in
- let nz96 : bits 4 := subrange_vec_dec v__438 7 4 in
- let nz54 : bits 2 := subrange_vec_dec v__438 3 2 in
- let nz3 : bits 1 := subrange_vec_dec v__438 1 1 in
- let nz2 : bits 1 := subrange_vec_dec v__438 0 0 in
+ | C_NOP tt =>
+ returnm (concat_vec ('b"000" : mword 3)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"00000" : mword 5) ('b"01" : mword 2)))))
+ | C_ADDI4SPN (rd, v__910) =>
+ (if let nz96 : bits 4 := subrange_vec_dec v__910 7 4 in
+ let nz96 : bits 4 := subrange_vec_dec v__910 7 4 in
+ let nz54 : bits 2 := subrange_vec_dec v__910 3 2 in
+ let nz3 : bits 1 := subrange_vec_dec v__910 1 1 in
+ let nz2 : bits 1 := subrange_vec_dec v__910 0 0 in
neq_vec (concat_vec nz96 (concat_vec nz54 (concat_vec nz3 nz2)))
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword (4 + (2 + (1 + 1)))))) then
- let nz96 : bits 4 := subrange_vec_dec v__438 7 4 in
- let nz96 : bits 4 := subrange_vec_dec v__438 7 4 in
- let nz54 : bits 2 := subrange_vec_dec v__438 3 2 in
- let nz3 : bits 1 := subrange_vec_dec v__438 1 1 in
- let nz2 : bits 1 := subrange_vec_dec v__438 0 0 in
- returnm ((concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (nz54 : bits 2)
- (concat_vec (nz96 : bits 4)
- (concat_vec (nz2 : bits 1)
- (concat_vec (nz3 : bits 1)
- (concat_vec (rd : cregidx) (vec_of_bits [B0;B0] : mword 2)))))))
- : mword 16)
+ (Ox"00"
+ : mword (4 + (2 + (1 + 1)))) then
+ let nz96 : bits 4 := subrange_vec_dec v__910 7 4 in
+ let nz96 : bits 4 := subrange_vec_dec v__910 7 4 in
+ let nz54 : bits 2 := subrange_vec_dec v__910 3 2 in
+ let nz3 : bits 1 := subrange_vec_dec v__910 1 1 in
+ let nz2 : bits 1 := subrange_vec_dec v__910 0 0 in
+ returnm (concat_vec ('b"000" : mword 3)
+ (concat_vec (nz54 : bits 2)
+ (concat_vec (nz96 : bits 4)
+ (concat_vec (nz2 : bits 1)
+ (concat_vec (nz3 : bits 1)
+ (concat_vec (rd : cregidx) ('b"00" : mword 2)))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_LW ((v__439, rs1, rd)) =>
- let ui6 : bits 1 := subrange_vec_dec v__439 4 4 in
- let ui6 : bits 1 := subrange_vec_dec v__439 4 4 in
- let ui53 : bits 3 := subrange_vec_dec v__439 3 1 in
- let ui2 : bits 1 := subrange_vec_dec v__439 0 0 in
- returnm ((concat_vec (vec_of_bits [B0;B1;B0] : mword 3)
- (concat_vec (ui53 : bits 3)
- (concat_vec (rs1 : cregidx)
- (concat_vec (ui2 : bits 1)
- (concat_vec (ui6 : bits 1)
- (concat_vec (rd : cregidx) (vec_of_bits [B0;B0] : mword 2)))))))
- : mword 16)
- | C_LD ((v__440, rs1, rd)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- let ui76 : bits 2 := subrange_vec_dec v__440 4 3 in
- let ui76 : bits 2 := subrange_vec_dec v__440 4 3 in
- let ui53 : bits 3 := subrange_vec_dec v__440 2 0 in
- returnm ((concat_vec (vec_of_bits [B0;B1;B1] : mword 3)
- (concat_vec (ui53 : bits 3)
- (concat_vec (rs1 : cregidx)
- (concat_vec (ui76 : bits 2)
- (concat_vec (rd : cregidx) (vec_of_bits [B0;B0] : mword 2))))))
- : mword 16)
+ | C_LW (v__911, rs1, rd) =>
+ let ui6 : bits 1 := subrange_vec_dec v__911 4 4 in
+ let ui6 : bits 1 := subrange_vec_dec v__911 4 4 in
+ let ui53 : bits 3 := subrange_vec_dec v__911 3 1 in
+ let ui2 : bits 1 := subrange_vec_dec v__911 0 0 in
+ returnm (concat_vec ('b"010" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (rs1 : cregidx)
+ (concat_vec (ui2 : bits 1)
+ (concat_vec (ui6 : bits 1)
+ (concat_vec (rd : cregidx) ('b"00" : mword 2)))))))
+ | C_LD (v__912, rs1, rd) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ let ui76 : bits 2 := subrange_vec_dec v__912 4 3 in
+ let ui76 : bits 2 := subrange_vec_dec v__912 4 3 in
+ let ui53 : bits 3 := subrange_vec_dec v__912 2 0 in
+ returnm (concat_vec ('b"011" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (rs1 : cregidx)
+ (concat_vec (ui76 : bits 2)
+ (concat_vec (rd : cregidx) ('b"00" : mword 2))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_SW ((v__441, rs1, rs2)) =>
- let ui6 : bits 1 := subrange_vec_dec v__441 4 4 in
- let ui6 : bits 1 := subrange_vec_dec v__441 4 4 in
- let ui53 : bits 3 := subrange_vec_dec v__441 3 1 in
- let ui2 : bits 1 := subrange_vec_dec v__441 0 0 in
- returnm ((concat_vec (vec_of_bits [B1;B1;B0] : mword 3)
- (concat_vec (ui53 : bits 3)
- (concat_vec (rs1 : cregidx)
- (concat_vec (ui2 : bits 1)
- (concat_vec (ui6 : bits 1)
- (concat_vec (rs2 : cregidx) (vec_of_bits [B0;B0] : mword 2)))))))
- : mword 16)
- | C_SD ((v__442, rs1, rs2)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- let ui76 : bits 2 := subrange_vec_dec v__442 4 3 in
- let ui76 : bits 2 := subrange_vec_dec v__442 4 3 in
- let ui53 : bits 3 := subrange_vec_dec v__442 2 0 in
- returnm ((concat_vec (vec_of_bits [B1;B1;B1] : mword 3)
- (concat_vec (ui53 : bits 3)
- (concat_vec (rs1 : bits 3)
- (concat_vec (ui76 : bits 2)
- (concat_vec (rs2 : bits 3) (vec_of_bits [B0;B0] : mword 2))))))
- : mword 16)
+ | C_SW (v__913, rs1, rs2) =>
+ let ui6 : bits 1 := subrange_vec_dec v__913 4 4 in
+ let ui6 : bits 1 := subrange_vec_dec v__913 4 4 in
+ let ui53 : bits 3 := subrange_vec_dec v__913 3 1 in
+ let ui2 : bits 1 := subrange_vec_dec v__913 0 0 in
+ returnm (concat_vec ('b"110" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (rs1 : cregidx)
+ (concat_vec (ui2 : bits 1)
+ (concat_vec (ui6 : bits 1)
+ (concat_vec (rs2 : cregidx) ('b"00" : mword 2)))))))
+ | C_SD (v__914, rs1, rs2) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ let ui76 : bits 2 := subrange_vec_dec v__914 4 3 in
+ let ui76 : bits 2 := subrange_vec_dec v__914 4 3 in
+ let ui53 : bits 3 := subrange_vec_dec v__914 2 0 in
+ returnm (concat_vec ('b"111" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (rs1 : bits 3)
+ (concat_vec (ui76 : bits 2)
+ (concat_vec (rs2 : bits 3) ('b"00" : mword 2))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_ADDI ((v__443, rsd)) =>
- (if ((let nzi5 : bits 1 := subrange_vec_dec v__443 5 5 in
- let nzi5 : bits 1 := subrange_vec_dec v__443 5 5 in
- let nzi40 : bits 5 := subrange_vec_dec v__443 4 0 in
- andb (neq_vec (concat_vec nzi5 nzi40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- : bool))) then
- let nzi5 : bits 1 := subrange_vec_dec v__443 5 5 in
- let nzi5 : bits 1 := subrange_vec_dec v__443 5 5 in
- let nzi40 : bits 5 := subrange_vec_dec v__443 4 0 in
- returnm ((concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (nzi5 : bits 1)
- (concat_vec (rsd : regidx)
- (concat_vec (nzi40 : bits 5) (vec_of_bits [B0;B1] : mword 2)))))
- : mword 16)
+ | C_ADDI (v__915, rsd) =>
+ (if let nzi5 : bits 1 := subrange_vec_dec v__915 5 5 in
+ let nzi5 : bits 1 := subrange_vec_dec v__915 5 5 in
+ let nzi40 : bits 5 := subrange_vec_dec v__915 4 0 in
+ andb (neq_vec (concat_vec nzi5 nzi40) ('b"000000" : mword (1 + 5))) (neq_vec rsd zreg)
+ then
+ let nzi5 : bits 1 := subrange_vec_dec v__915 5 5 in
+ let nzi5 : bits 1 := subrange_vec_dec v__915 5 5 in
+ let nzi40 : bits 5 := subrange_vec_dec v__915 4 0 in
+ returnm (concat_vec ('b"000" : mword 3)
+ (concat_vec (nzi5 : bits 1)
+ (concat_vec (rsd : regidx)
+ (concat_vec (nzi40 : bits 5) ('b"01" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_JAL (v__444) =>
- (if sumbool_of_bool ((Z.eqb 64 32)) then
- let i11 : bits 1 := subrange_vec_dec v__444 10 10 in
- let i98 : bits 2 := subrange_vec_dec v__444 8 7 in
- let i7 : bits 1 := subrange_vec_dec v__444 6 6 in
- let i6 : bits 1 := subrange_vec_dec v__444 5 5 in
- let i5 : bits 1 := subrange_vec_dec v__444 4 4 in
- let i4 : bits 1 := subrange_vec_dec v__444 3 3 in
- let i31 : bits 3 := subrange_vec_dec v__444 2 0 in
- let i11 : bits 1 := subrange_vec_dec v__444 10 10 in
- let i10 : bits 1 := subrange_vec_dec v__444 9 9 in
- returnm ((concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
- (concat_vec (i11 : bits 1)
- (concat_vec (i4 : bits 1)
- (concat_vec (i98 : bits 2)
- (concat_vec (i10 : bits 1)
- (concat_vec (i6 : bits 1)
- (concat_vec (i7 : bits 1)
- (concat_vec (i31 : bits 3)
- (concat_vec (i5 : bits 1)
- (vec_of_bits [B0;B1] : mword 2))))))))))
- : mword 16)
+ | C_JAL v__916 =>
+ (if sumbool_of_bool (Z.eqb 64 32) then
+ let i11 : bits 1 := subrange_vec_dec v__916 10 10 in
+ let i98 : bits 2 := subrange_vec_dec v__916 8 7 in
+ let i7 : bits 1 := subrange_vec_dec v__916 6 6 in
+ let i6 : bits 1 := subrange_vec_dec v__916 5 5 in
+ let i5 : bits 1 := subrange_vec_dec v__916 4 4 in
+ let i4 : bits 1 := subrange_vec_dec v__916 3 3 in
+ let i31 : bits 3 := subrange_vec_dec v__916 2 0 in
+ let i11 : bits 1 := subrange_vec_dec v__916 10 10 in
+ let i10 : bits 1 := subrange_vec_dec v__916 9 9 in
+ returnm (concat_vec ('b"001" : mword 3)
+ (concat_vec (i11 : bits 1)
+ (concat_vec (i4 : bits 1)
+ (concat_vec (i98 : bits 2)
+ (concat_vec (i10 : bits 1)
+ (concat_vec (i6 : bits 1)
+ (concat_vec (i7 : bits 1)
+ (concat_vec (i31 : bits 3)
+ (concat_vec (i5 : bits 1) ('b"01" : mword 2))))))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_ADDIW ((v__445, rsd)) =>
- (if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
- (projT1
- (regidx_to_regno zreg)))) (Z.eqb 64 64))) then
- let imm5 : bits 1 := subrange_vec_dec v__445 5 5 in
- let imm5 : bits 1 := subrange_vec_dec v__445 5 5 in
- let imm40 : bits 5 := subrange_vec_dec v__445 4 0 in
- returnm ((concat_vec (vec_of_bits [B0;B0;B1] : mword 3)
- (concat_vec (imm5 : bits 1)
- (concat_vec (rsd : regidx)
- (concat_vec (imm40 : bits 5) (vec_of_bits [B0;B1] : mword 2)))))
- : mword 16)
+ | C_ADDIW (v__917, rsd) =>
+ (if sumbool_of_bool (andb (neq_vec rsd zreg) (Z.eqb 64 64)) then
+ let imm5 : bits 1 := subrange_vec_dec v__917 5 5 in
+ let imm5 : bits 1 := subrange_vec_dec v__917 5 5 in
+ let imm40 : bits 5 := subrange_vec_dec v__917 4 0 in
+ returnm (concat_vec ('b"001" : mword 3)
+ (concat_vec (imm5 : bits 1)
+ (concat_vec (rsd : regidx)
+ (concat_vec (imm40 : bits 5) ('b"01" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_LI ((v__446, rd)) =>
- (if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg)))))
- then
- let imm5 : bits 1 := subrange_vec_dec v__446 5 5 in
- let imm5 : bits 1 := subrange_vec_dec v__446 5 5 in
- let imm40 : bits 5 := subrange_vec_dec v__446 4 0 in
- returnm ((concat_vec (vec_of_bits [B0;B1;B0] : mword 3)
- (concat_vec (imm5 : bits 1)
- (concat_vec (rd : regidx)
- (concat_vec (imm40 : bits 5) (vec_of_bits [B0;B1] : mword 2)))))
- : mword 16)
+ | C_LI (v__918, rd) =>
+ (if neq_vec rd zreg then
+ let imm5 : bits 1 := subrange_vec_dec v__918 5 5 in
+ let imm5 : bits 1 := subrange_vec_dec v__918 5 5 in
+ let imm40 : bits 5 := subrange_vec_dec v__918 4 0 in
+ returnm (concat_vec ('b"010" : mword 3)
+ (concat_vec (imm5 : bits 1)
+ (concat_vec (rd : regidx)
+ (concat_vec (imm40 : bits 5) ('b"01" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_ADDI16SP (v__447) =>
- (if ((let nzi9 : bits 1 := subrange_vec_dec v__447 5 5 in
- let nzi9 : bits 1 := subrange_vec_dec v__447 5 5 in
- let nzi87 : bits 2 := subrange_vec_dec v__447 4 3 in
- let nzi6 : bits 1 := subrange_vec_dec v__447 2 2 in
- let nzi5 : bits 1 := subrange_vec_dec v__447 1 1 in
- let nzi4 : bits 1 := subrange_vec_dec v__447 0 0 in
+ | C_ADDI16SP v__919 =>
+ (if let nzi9 : bits 1 := subrange_vec_dec v__919 5 5 in
+ let nzi9 : bits 1 := subrange_vec_dec v__919 5 5 in
+ let nzi87 : bits 2 := subrange_vec_dec v__919 4 3 in
+ let nzi6 : bits 1 := subrange_vec_dec v__919 2 2 in
+ let nzi5 : bits 1 := subrange_vec_dec v__919 1 1 in
+ let nzi4 : bits 1 := subrange_vec_dec v__919 0 0 in
neq_vec (concat_vec nzi9 (concat_vec nzi87 (concat_vec nzi6 (concat_vec nzi5 nzi4))))
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + (2 + (1 + (1 + 1))))))) then
- let nzi9 : bits 1 := subrange_vec_dec v__447 5 5 in
- let nzi9 : bits 1 := subrange_vec_dec v__447 5 5 in
- let nzi87 : bits 2 := subrange_vec_dec v__447 4 3 in
- let nzi6 : bits 1 := subrange_vec_dec v__447 2 2 in
- let nzi5 : bits 1 := subrange_vec_dec v__447 1 1 in
- let nzi4 : bits 1 := subrange_vec_dec v__447 0 0 in
- returnm ((concat_vec (vec_of_bits [B0;B1;B1] : mword 3)
- (concat_vec (nzi9 : bits 1)
- (concat_vec (vec_of_bits [B0;B0;B0;B1;B0] : mword 5)
- (concat_vec (nzi4 : bits 1)
- (concat_vec (nzi6 : bits 1)
- (concat_vec (nzi87 : bits 2)
- (concat_vec (nzi5 : bits 1) (vec_of_bits [B0;B1] : mword 2))))))))
- : mword 16)
+ ('b"000000"
+ : mword (1 + (2 + (1 + (1 + 1))))) then
+ let nzi9 : bits 1 := subrange_vec_dec v__919 5 5 in
+ let nzi9 : bits 1 := subrange_vec_dec v__919 5 5 in
+ let nzi87 : bits 2 := subrange_vec_dec v__919 4 3 in
+ let nzi6 : bits 1 := subrange_vec_dec v__919 2 2 in
+ let nzi5 : bits 1 := subrange_vec_dec v__919 1 1 in
+ let nzi4 : bits 1 := subrange_vec_dec v__919 0 0 in
+ returnm (concat_vec ('b"011" : mword 3)
+ (concat_vec (nzi9 : bits 1)
+ (concat_vec ('b"00010" : mword 5)
+ (concat_vec (nzi4 : bits 1)
+ (concat_vec (nzi6 : bits 1)
+ (concat_vec (nzi87 : bits 2)
+ (concat_vec (nzi5 : bits 1) ('b"01" : mword 2))))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_LUI ((v__448, rd)) =>
- (if sumbool_of_bool ((let imm17 : bits 1 := subrange_vec_dec v__448 5 5 in
- let imm17 : bits 1 := subrange_vec_dec v__448 5 5 in
- let imm1612 : bits 5 := subrange_vec_dec v__448 4 0 in
- andb (projT1 (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg))))
- ((andb (projT1 (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno sp))))
- (neq_vec (concat_vec imm17 imm1612)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5))))
- : bool))) then
- let imm17 : bits 1 := subrange_vec_dec v__448 5 5 in
- let imm17 : bits 1 := subrange_vec_dec v__448 5 5 in
- let imm1612 : bits 5 := subrange_vec_dec v__448 4 0 in
- returnm ((concat_vec (vec_of_bits [B0;B1;B1] : mword 3)
- (concat_vec (imm17 : bits 1)
- (concat_vec (rd : regidx)
- (concat_vec (imm1612 : bits 5) (vec_of_bits [B0;B1] : mword 2)))))
- : mword 16)
+ | C_LUI (v__920, rd) =>
+ (if let imm17 : bits 1 := subrange_vec_dec v__920 5 5 in
+ let imm17 : bits 1 := subrange_vec_dec v__920 5 5 in
+ let imm1612 : bits 5 := subrange_vec_dec v__920 4 0 in
+ andb (neq_vec rd zreg)
+ (andb (neq_vec rd sp) (neq_vec (concat_vec imm17 imm1612) ('b"000000" : mword (1 + 5))))
+ then
+ let imm17 : bits 1 := subrange_vec_dec v__920 5 5 in
+ let imm17 : bits 1 := subrange_vec_dec v__920 5 5 in
+ let imm1612 : bits 5 := subrange_vec_dec v__920 4 0 in
+ returnm (concat_vec ('b"011" : mword 3)
+ (concat_vec (imm17 : bits 1)
+ (concat_vec (rd : regidx)
+ (concat_vec (imm1612 : bits 5) ('b"01" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_SRLI ((v__449, rsd)) =>
- (if ((let nzui5 : bits 1 := subrange_vec_dec v__449 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__449 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__449 4 0 in
- neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5))))
- then
- let nzui5 : bits 1 := subrange_vec_dec v__449 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__449 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__449 4 0 in
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (nzui5 : bits 1)
- (concat_vec (vec_of_bits [B0;B0] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (nzui40 : bits 5) (vec_of_bits [B0;B1] : mword 2))))))
- : mword 16)
+ | C_SRLI (v__921, rsd) =>
+ (if let nzui5 : bits 1 := subrange_vec_dec v__921 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__921 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__921 4 0 in
+ neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)) then
+ let nzui5 : bits 1 := subrange_vec_dec v__921 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__921 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__921 4 0 in
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec (nzui5 : bits 1)
+ (concat_vec ('b"00" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec (nzui40 : bits 5) ('b"01" : mword 2))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_SRAI ((v__450, rsd)) =>
- (if ((let nzui5 : bits 1 := subrange_vec_dec v__450 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__450 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__450 4 0 in
- neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5))))
- then
- let nzui5 : bits 1 := subrange_vec_dec v__450 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__450 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__450 4 0 in
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (nzui5 : bits 1)
- (concat_vec (vec_of_bits [B0;B1] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (nzui40 : bits 5) (vec_of_bits [B0;B1] : mword 2))))))
- : mword 16)
+ | C_SRAI (v__922, rsd) =>
+ (if let nzui5 : bits 1 := subrange_vec_dec v__922 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__922 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__922 4 0 in
+ neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)) then
+ let nzui5 : bits 1 := subrange_vec_dec v__922 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__922 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__922 4 0 in
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec (nzui5 : bits 1)
+ (concat_vec ('b"01" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec (nzui40 : bits 5) ('b"01" : mword 2))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_ANDI ((v__451, rsd)) =>
- let i5 : bits 1 := subrange_vec_dec v__451 5 5 in
- let i5 : bits 1 := subrange_vec_dec v__451 5 5 in
- let i40 : bits 5 := subrange_vec_dec v__451 4 0 in
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (i5 : bits 1)
- (concat_vec (vec_of_bits [B1;B0] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (i40 : bits 5) (vec_of_bits [B0;B1] : mword 2))))))
- : mword 16)
- | C_SUB ((rsd, rs2)) =>
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (vec_of_bits [B0;B0] : mword 2)
- (concat_vec (rs2 : cregidx) (vec_of_bits [B0;B1] : mword 2)))))))
- : mword (3 + (1 + (2 + (3 + (2 + (3 + 2)))))))
- | C_XOR ((rsd, rs2)) =>
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (vec_of_bits [B0;B1] : mword 2)
- (concat_vec (rs2 : cregidx) (vec_of_bits [B0;B1] : mword 2)))))))
- : mword (3 + (1 + (2 + (3 + (2 + (3 + 2)))))))
- | C_OR ((rsd, rs2)) =>
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (vec_of_bits [B1;B0] : mword 2)
- (concat_vec (rs2 : cregidx) (vec_of_bits [B0;B1] : mword 2)))))))
- : mword (3 + (1 + (2 + (3 + (2 + (3 + 2)))))))
- | C_AND ((rsd, rs2)) =>
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (rs2 : cregidx) (vec_of_bits [B0;B1] : mword 2)))))))
- : mword (3 + (1 + (2 + (3 + (2 + (3 + 2)))))))
- | C_SUBW ((rsd, rs2)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B1] : mword 1)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (vec_of_bits [B0;B0] : mword 2)
- (concat_vec (rs2 : cregidx) (vec_of_bits [B0;B1] : mword 2)))))))
- : mword (3 + (1 + (2 + (3 + (2 + (3 + 2)))))))
+ | C_ANDI (v__923, rsd) =>
+ let i5 : bits 1 := subrange_vec_dec v__923 5 5 in
+ let i5 : bits 1 := subrange_vec_dec v__923 5 5 in
+ let i40 : bits 5 := subrange_vec_dec v__923 4 0 in
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec (i5 : bits 1)
+ (concat_vec ('b"10" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec (i40 : bits 5) ('b"01" : mword 2))))))
+ | C_SUB (rsd, rs2) =>
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec ('b"00" : mword 2)
+ (concat_vec (rs2 : cregidx) ('b"01" : mword 2)))))))
+ | C_XOR (rsd, rs2) =>
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec ('b"01" : mword 2)
+ (concat_vec (rs2 : cregidx) ('b"01" : mword 2)))))))
+ | C_OR (rsd, rs2) =>
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec ('b"10" : mword 2)
+ (concat_vec (rs2 : cregidx) ('b"01" : mword 2)))))))
+ | C_AND (rsd, rs2) =>
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (rs2 : cregidx) ('b"01" : mword 2)))))))
+ | C_SUBW (rsd, rs2) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"1" : mword 1)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec ('b"00" : mword 2)
+ (concat_vec (rs2 : cregidx) ('b"01" : mword 2)))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_ADDW ((rsd, rs2)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B1] : mword 1)
- (concat_vec (vec_of_bits [B1;B1] : mword 2)
- (concat_vec (rsd : cregidx)
- (concat_vec (vec_of_bits [B0;B1] : mword 2)
- (concat_vec (rs2 : cregidx) (vec_of_bits [B0;B1] : mword 2)))))))
- : mword (3 + (1 + (2 + (3 + (2 + (3 + 2)))))))
+ | C_ADDW (rsd, rs2) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"1" : mword 1)
+ (concat_vec ('b"11" : mword 2)
+ (concat_vec (rsd : cregidx)
+ (concat_vec ('b"01" : mword 2)
+ (concat_vec (rs2 : cregidx) ('b"01" : mword 2)))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_J (v__452) =>
- let i11 : bits 1 := subrange_vec_dec v__452 10 10 in
- let i98 : bits 2 := subrange_vec_dec v__452 8 7 in
- let i7 : bits 1 := subrange_vec_dec v__452 6 6 in
- let i6 : bits 1 := subrange_vec_dec v__452 5 5 in
- let i5 : bits 1 := subrange_vec_dec v__452 4 4 in
- let i4 : bits 1 := subrange_vec_dec v__452 3 3 in
- let i31 : bits 3 := subrange_vec_dec v__452 2 0 in
- let i11 : bits 1 := subrange_vec_dec v__452 10 10 in
- let i10 : bits 1 := subrange_vec_dec v__452 9 9 in
- returnm ((concat_vec (vec_of_bits [B1;B0;B1] : mword 3)
- (concat_vec (i11 : bits 1)
- (concat_vec (i4 : bits 1)
- (concat_vec (i98 : bits 2)
- (concat_vec (i10 : bits 1)
- (concat_vec (i6 : bits 1)
- (concat_vec (i7 : bits 1)
- (concat_vec (i31 : bits 3)
- (concat_vec (i5 : bits 1) (vec_of_bits [B0;B1] : mword 2))))))))))
- : mword 16)
- | C_BEQZ ((v__453, rs)) =>
- let i8 : bits 1 := subrange_vec_dec v__453 7 7 in
- let i8 : bits 1 := subrange_vec_dec v__453 7 7 in
- let i76 : bits 2 := subrange_vec_dec v__453 6 5 in
- let i5 : bits 1 := subrange_vec_dec v__453 4 4 in
- let i43 : bits 2 := subrange_vec_dec v__453 3 2 in
- let i21 : bits 2 := subrange_vec_dec v__453 1 0 in
- returnm ((concat_vec (vec_of_bits [B1;B1;B0] : mword 3)
- (concat_vec (i8 : bits 1)
- (concat_vec (i43 : bits 2)
- (concat_vec (rs : cregidx)
- (concat_vec (i76 : bits 2)
- (concat_vec (i21 : bits 2)
- (concat_vec (i5 : bits 1) (vec_of_bits [B0;B1] : mword 2))))))))
- : mword 16)
- | C_BNEZ ((v__454, rs)) =>
- let i8 : bits 1 := subrange_vec_dec v__454 7 7 in
- let i8 : bits 1 := subrange_vec_dec v__454 7 7 in
- let i76 : bits 2 := subrange_vec_dec v__454 6 5 in
- let i5 : bits 1 := subrange_vec_dec v__454 4 4 in
- let i43 : bits 2 := subrange_vec_dec v__454 3 2 in
- let i21 : bits 2 := subrange_vec_dec v__454 1 0 in
- returnm ((concat_vec (vec_of_bits [B1;B1;B1] : mword 3)
- (concat_vec (i8 : bits 1)
- (concat_vec (i43 : bits 2)
- (concat_vec (rs : cregidx)
- (concat_vec (i76 : bits 2)
- (concat_vec (i21 : bits 2)
- (concat_vec (i5 : bits 1) (vec_of_bits [B0;B1] : mword 2))))))))
- : mword 16)
- | C_SLLI ((v__455, rsd)) =>
- (if ((let nzui5 : bits 1 := subrange_vec_dec v__455 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__455 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__455 4 0 in
- andb
- (neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- ((andb (projT1 (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- (orb (Z.eqb 64 64) (eq_vec nzui5 ((bool_to_bits false) : mword 1))))
- : bool))) then
- let nzui5 : bits 1 := subrange_vec_dec v__455 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__455 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__455 4 0 in
- returnm ((concat_vec (vec_of_bits [B0;B0;B0] : mword 3)
- (concat_vec (nzui5 : bits 1)
- (concat_vec (rsd : regidx)
- (concat_vec (nzui40 : bits 5) (vec_of_bits [B1;B0] : mword 2)))))
- : mword 16)
+ | C_J v__924 =>
+ let i11 : bits 1 := subrange_vec_dec v__924 10 10 in
+ let i98 : bits 2 := subrange_vec_dec v__924 8 7 in
+ let i7 : bits 1 := subrange_vec_dec v__924 6 6 in
+ let i6 : bits 1 := subrange_vec_dec v__924 5 5 in
+ let i5 : bits 1 := subrange_vec_dec v__924 4 4 in
+ let i4 : bits 1 := subrange_vec_dec v__924 3 3 in
+ let i31 : bits 3 := subrange_vec_dec v__924 2 0 in
+ let i11 : bits 1 := subrange_vec_dec v__924 10 10 in
+ let i10 : bits 1 := subrange_vec_dec v__924 9 9 in
+ returnm (concat_vec ('b"101" : mword 3)
+ (concat_vec (i11 : bits 1)
+ (concat_vec (i4 : bits 1)
+ (concat_vec (i98 : bits 2)
+ (concat_vec (i10 : bits 1)
+ (concat_vec (i6 : bits 1)
+ (concat_vec (i7 : bits 1)
+ (concat_vec (i31 : bits 3)
+ (concat_vec (i5 : bits 1) ('b"01" : mword 2))))))))))
+ | C_BEQZ (v__925, rs) =>
+ let i8 : bits 1 := subrange_vec_dec v__925 7 7 in
+ let i8 : bits 1 := subrange_vec_dec v__925 7 7 in
+ let i76 : bits 2 := subrange_vec_dec v__925 6 5 in
+ let i5 : bits 1 := subrange_vec_dec v__925 4 4 in
+ let i43 : bits 2 := subrange_vec_dec v__925 3 2 in
+ let i21 : bits 2 := subrange_vec_dec v__925 1 0 in
+ returnm (concat_vec ('b"110" : mword 3)
+ (concat_vec (i8 : bits 1)
+ (concat_vec (i43 : bits 2)
+ (concat_vec (rs : cregidx)
+ (concat_vec (i76 : bits 2)
+ (concat_vec (i21 : bits 2)
+ (concat_vec (i5 : bits 1) ('b"01" : mword 2))))))))
+ | C_BNEZ (v__926, rs) =>
+ let i8 : bits 1 := subrange_vec_dec v__926 7 7 in
+ let i8 : bits 1 := subrange_vec_dec v__926 7 7 in
+ let i76 : bits 2 := subrange_vec_dec v__926 6 5 in
+ let i5 : bits 1 := subrange_vec_dec v__926 4 4 in
+ let i43 : bits 2 := subrange_vec_dec v__926 3 2 in
+ let i21 : bits 2 := subrange_vec_dec v__926 1 0 in
+ returnm (concat_vec ('b"111" : mword 3)
+ (concat_vec (i8 : bits 1)
+ (concat_vec (i43 : bits 2)
+ (concat_vec (rs : cregidx)
+ (concat_vec (i76 : bits 2)
+ (concat_vec (i21 : bits 2)
+ (concat_vec (i5 : bits 1) ('b"01" : mword 2))))))))
+ | C_SLLI (v__927, rsd) =>
+ (if sumbool_of_bool
+ (let nzui5 : bits 1 := subrange_vec_dec v__927 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__927 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__927 4 0 in
+ andb (neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)))
+ (andb (neq_vec rsd zreg) (orb (Z.eqb 64 64) (eq_vec nzui5 ('b"0" : mword 1))))) then
+ let nzui5 : bits 1 := subrange_vec_dec v__927 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__927 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__927 4 0 in
+ returnm (concat_vec ('b"000" : mword 3)
+ (concat_vec (nzui5 : bits 1)
+ (concat_vec (rsd : regidx)
+ (concat_vec (nzui40 : bits 5) ('b"10" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_LWSP ((v__456, rd)) =>
- (if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg)))))
- then
- let ui76 : bits 2 := subrange_vec_dec v__456 5 4 in
- let ui76 : bits 2 := subrange_vec_dec v__456 5 4 in
- let ui5 : bits 1 := subrange_vec_dec v__456 3 3 in
- let ui42 : bits 3 := subrange_vec_dec v__456 2 0 in
- returnm ((concat_vec (vec_of_bits [B0;B1;B0] : mword 3)
- (concat_vec (ui5 : bits 1)
- (concat_vec (rd : regidx)
- (concat_vec (ui42 : bits 3)
- (concat_vec (ui76 : bits 2) (vec_of_bits [B1;B0] : mword 2))))))
- : mword 16)
+ | C_LWSP (v__928, rd) =>
+ (if neq_vec rd zreg then
+ let ui76 : bits 2 := subrange_vec_dec v__928 5 4 in
+ let ui76 : bits 2 := subrange_vec_dec v__928 5 4 in
+ let ui5 : bits 1 := subrange_vec_dec v__928 3 3 in
+ let ui42 : bits 3 := subrange_vec_dec v__928 2 0 in
+ returnm (concat_vec ('b"010" : mword 3)
+ (concat_vec (ui5 : bits 1)
+ (concat_vec (rd : regidx)
+ (concat_vec (ui42 : bits 3)
+ (concat_vec (ui76 : bits 2) ('b"10" : mword 2))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_LDSP ((v__457, rd)) =>
- (if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg)))) (Z.eqb 64 64))) then
- let ui86 : bits 3 := subrange_vec_dec v__457 5 3 in
- let ui86 : bits 3 := subrange_vec_dec v__457 5 3 in
- let ui5 : bits 1 := subrange_vec_dec v__457 2 2 in
- let ui43 : bits 2 := subrange_vec_dec v__457 1 0 in
- returnm ((concat_vec (vec_of_bits [B0;B1;B1] : mword 3)
- (concat_vec (ui5 : bits 1)
- (concat_vec (rd : regidx)
- (concat_vec (ui43 : bits 2)
- (concat_vec (ui86 : bits 3) (vec_of_bits [B1;B0] : mword 2))))))
- : mword 16)
+ | C_LDSP (v__929, rd) =>
+ (if sumbool_of_bool (andb (neq_vec rd zreg) (Z.eqb 64 64)) then
+ let ui86 : bits 3 := subrange_vec_dec v__929 5 3 in
+ let ui86 : bits 3 := subrange_vec_dec v__929 5 3 in
+ let ui5 : bits 1 := subrange_vec_dec v__929 2 2 in
+ let ui43 : bits 2 := subrange_vec_dec v__929 1 0 in
+ returnm (concat_vec ('b"011" : mword 3)
+ (concat_vec (ui5 : bits 1)
+ (concat_vec (rd : regidx)
+ (concat_vec (ui43 : bits 2)
+ (concat_vec (ui86 : bits 3) ('b"10" : mword 2))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_SWSP ((v__458, rs2)) =>
- let ui76 : bits 2 := subrange_vec_dec v__458 5 4 in
- let ui76 : bits 2 := subrange_vec_dec v__458 5 4 in
- let ui52 : bits 4 := subrange_vec_dec v__458 3 0 in
- returnm ((concat_vec (vec_of_bits [B1;B1;B0] : mword 3)
- (concat_vec (ui52 : bits 4)
- (concat_vec (ui76 : bits 2)
- (concat_vec (rs2 : regidx) (vec_of_bits [B1;B0] : mword 2)))))
- : mword 16)
- | C_SDSP ((v__459, rs2)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
- let ui86 : bits 3 := subrange_vec_dec v__459 5 3 in
- let ui86 : bits 3 := subrange_vec_dec v__459 5 3 in
- let ui53 : bits 3 := subrange_vec_dec v__459 2 0 in
- returnm ((concat_vec (vec_of_bits [B1;B1;B1] : mword 3)
- (concat_vec (ui53 : bits 3)
- (concat_vec (ui86 : bits 3)
- (concat_vec (rs2 : regidx) (vec_of_bits [B1;B0] : mword 2)))))
- : mword 16)
+ | C_SWSP (v__930, rs2) =>
+ let ui76 : bits 2 := subrange_vec_dec v__930 5 4 in
+ let ui76 : bits 2 := subrange_vec_dec v__930 5 4 in
+ let ui52 : bits 4 := subrange_vec_dec v__930 3 0 in
+ returnm (concat_vec ('b"110" : mword 3)
+ (concat_vec (ui52 : bits 4)
+ (concat_vec (ui76 : bits 2) (concat_vec (rs2 : regidx) ('b"10" : mword 2)))))
+ | C_SDSP (v__931, rs2) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
+ let ui86 : bits 3 := subrange_vec_dec v__931 5 3 in
+ let ui86 : bits 3 := subrange_vec_dec v__931 5 3 in
+ let ui53 : bits 3 := subrange_vec_dec v__931 2 0 in
+ returnm (concat_vec ('b"111" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (ui86 : bits 3)
+ (concat_vec (rs2 : regidx) ('b"10" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_JR (rs1) =>
- (if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rs1)) (projT1 (regidx_to_regno zreg)))))
- then
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (rs1 : regidx)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B0] : mword 2)))))
- : mword (3 + (1 + (5 + (5 + 2)))))
+ | C_JR rs1 =>
+ (if neq_vec rs1 zreg then
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec (rs1 : regidx)
+ (concat_vec ('b"00000" : mword 5) ('b"10" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_JALR (rs1) =>
- (if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rs1)) (projT1 (regidx_to_regno zreg)))))
- then
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B1] : mword 1)
- (concat_vec (rs1 : regidx)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B0] : mword 2)))))
- : mword (3 + (1 + (5 + (5 + 2)))))
+ | C_JALR rs1 =>
+ (if neq_vec rs1 zreg then
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"1" : mword 1)
+ (concat_vec (rs1 : regidx)
+ (concat_vec ('b"00000" : mword 5) ('b"10" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_MV ((rd, rs2)) =>
- (if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))) then
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B0] : mword 1)
- (concat_vec (rd : regidx)
- (concat_vec (rs2 : regidx) (vec_of_bits [B1;B0] : mword 2)))))
- : mword (3 + (1 + (5 + (5 + 2)))))
+ | C_MV (rd, rs2) =>
+ (if andb (neq_vec rd zreg) (neq_vec rs2 zreg) then
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"0" : mword 1)
+ (concat_vec (rd : regidx) (concat_vec (rs2 : regidx) ('b"10" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_EBREAK (tt) =>
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B1] : mword 1)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (concat_vec (vec_of_bits [B0;B0;B0;B0;B0] : mword 5)
- (vec_of_bits [B1;B0] : mword 2)))))
- : mword (3 + (1 + (5 + (5 + 2)))))
- | C_ADD ((rsd, rs2)) =>
- (if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
- (projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))) then
- returnm ((concat_vec (vec_of_bits [B1;B0;B0] : mword 3)
- (concat_vec (vec_of_bits [B1] : mword 1)
- (concat_vec (rsd : regidx)
- (concat_vec (rs2 : regidx) (vec_of_bits [B1;B0] : mword 2)))))
- : mword (3 + (1 + (5 + (5 + 2)))))
+ | C_EBREAK tt =>
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"1" : mword 1)
+ (concat_vec ('b"00000" : mword 5)
+ (concat_vec ('b"00000" : mword 5) ('b"10" : mword 2)))))
+ | C_ADD (rsd, rs2) =>
+ (if andb (neq_vec rsd zreg) (neq_vec rs2 zreg) then
+ returnm (concat_vec ('b"100" : mword 3)
+ (concat_vec ('b"1" : mword 1)
+ (concat_vec (rsd : regidx) (concat_vec (rs2 : regidx) ('b"10" : mword 2)))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 16)
+ | C_FLWSP (v__932, rd) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__45 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__45 then
+ let ui76 : bits 2 := subrange_vec_dec v__932 5 4 in
+ let ui76 : bits 2 := subrange_vec_dec v__932 5 4 in
+ let ui5 : bits 1 := subrange_vec_dec v__932 3 3 in
+ let ui42 : bits 3 := subrange_vec_dec v__932 2 0 in
+ returnm (concat_vec ('b"011" : mword 3)
+ (concat_vec (ui5 : bits 1)
+ (concat_vec (rd : regidx)
+ (concat_vec (ui42 : bits 3)
+ (concat_vec (ui76 : bits 2) ('b"10" : mword 2))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 16)
+ | C_FSWSP (v__933, rs2) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__51 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__51 then
+ let ui76 : bits 2 := subrange_vec_dec v__933 5 4 in
+ let ui76 : bits 2 := subrange_vec_dec v__933 5 4 in
+ let ui52 : bits 4 := subrange_vec_dec v__933 3 0 in
+ returnm (concat_vec ('b"111" : mword 3)
+ (concat_vec (ui52 : bits 4)
+ (concat_vec (ui76 : bits 2)
+ (concat_vec (rs2 : regidx) ('b"10" : mword 2)))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (mword 16)
- | C_ILLEGAL (s) => returnm (s : mword 16)
+ | C_FLW (v__934, rs1, rd) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__57 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__57 then
+ let ui6 : bits 1 := subrange_vec_dec v__934 4 4 in
+ let ui6 : bits 1 := subrange_vec_dec v__934 4 4 in
+ let ui53 : bits 3 := subrange_vec_dec v__934 3 1 in
+ let ui2 : bits 1 := subrange_vec_dec v__934 0 0 in
+ returnm (concat_vec ('b"011" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (rs1 : cregidx)
+ (concat_vec (ui2 : bits 1)
+ (concat_vec (ui6 : bits 1)
+ (concat_vec (rd : cregidx) ('b"00" : mword 2)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 16)
+ | C_FSW (v__935, rs1, rs2) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__63 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__63 then
+ let ui6 : bits 1 := subrange_vec_dec v__935 4 4 in
+ let ui6 : bits 1 := subrange_vec_dec v__935 4 4 in
+ let ui53 : bits 3 := subrange_vec_dec v__935 3 1 in
+ let ui2 : bits 1 := subrange_vec_dec v__935 0 0 in
+ returnm (concat_vec ('b"111" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (rs1 : cregidx)
+ (concat_vec (ui2 : bits 1)
+ (concat_vec (ui6 : bits 1)
+ (concat_vec (rs2 : cregidx) ('b"00" : mword 2)))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 16)
+ | C_FLDSP (v__936, rd) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__69 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__69 then
+ let ui86 : bits 3 := subrange_vec_dec v__936 5 3 in
+ let ui86 : bits 3 := subrange_vec_dec v__936 5 3 in
+ let ui5 : bits 1 := subrange_vec_dec v__936 2 2 in
+ let ui43 : bits 2 := subrange_vec_dec v__936 1 0 in
+ returnm (concat_vec ('b"001" : mword 3)
+ (concat_vec (ui5 : bits 1)
+ (concat_vec (rd : regidx)
+ (concat_vec (ui43 : bits 2)
+ (concat_vec (ui86 : bits 3) ('b"10" : mword 2))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 16)
+ | C_FSDSP (v__937, rs2) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__75 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__75 then
+ let ui86 : bits 3 := subrange_vec_dec v__937 5 3 in
+ let ui86 : bits 3 := subrange_vec_dec v__937 5 3 in
+ let ui53 : bits 3 := subrange_vec_dec v__937 2 0 in
+ returnm (concat_vec ('b"101" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (ui86 : bits 3)
+ (concat_vec (rs2 : regidx) ('b"10" : mword 2)))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 16)
+ | C_FLD (v__938, rs1, rd) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__81 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__81 then
+ let ui76 : bits 2 := subrange_vec_dec v__938 4 3 in
+ let ui76 : bits 2 := subrange_vec_dec v__938 4 3 in
+ let ui53 : bits 3 := subrange_vec_dec v__938 2 0 in
+ returnm (concat_vec ('b"001" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (rs1 : cregidx)
+ (concat_vec (ui76 : bits 2)
+ (concat_vec (rd : cregidx) ('b"00" : mword 2))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 16)
+ | C_FSD (v__939, rs1, rs2) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__87 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__87 then
+ let ui76 : bits 2 := subrange_vec_dec v__939 4 3 in
+ let ui76 : bits 2 := subrange_vec_dec v__939 4 3 in
+ let ui53 : bits 3 := subrange_vec_dec v__939 2 0 in
+ returnm (concat_vec ('b"101" : mword 3)
+ (concat_vec (ui53 : bits 3)
+ (concat_vec (rs1 : bits 3)
+ (concat_vec (ui76 : bits 2)
+ (concat_vec (rs2 : bits 3) ('b"00" : mword 2))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (mword 16)
+ | C_ILLEGAL s => returnm s
| _ => assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt
end)
: M (mword 16).
-Definition encdec_compressed_backwards (arg_ : mword 16)
-: ast :=
-
- let v__460 := arg_ in
- if ((eq_vec v__460 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 16)))
- then
- C_NOP
- (tt)
- else if ((andb
- (let nz96 : bits 4 := subrange_vec_dec v__460 10 7 in
- let nz54 : bits 2 := subrange_vec_dec v__460 12 11 in
- let nz3 : bits 1 := subrange_vec_dec v__460 5 5 in
- let nz2 : bits 1 := subrange_vec_dec v__460 6 6 in
- neq_vec (concat_vec nz96 (concat_vec nz54 (concat_vec nz3 nz2)))
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword (4 + (2 + (1 + 1)))))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B0;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B0] : mword (1 - 0 + 1))))))
- then
- let rd : cregidx := subrange_vec_dec v__460 4 2 in
- let nz96 : bits 4 := subrange_vec_dec v__460 10 7 in
- let nz54 : bits 2 := subrange_vec_dec v__460 12 11 in
- let nz3 : bits 1 := subrange_vec_dec v__460 5 5 in
- let nz2 : bits 1 := subrange_vec_dec v__460 6 6 in
- C_ADDI4SPN
- ((rd, concat_vec (nz96 : bits 4)
- (concat_vec (nz54 : bits 2) (concat_vec (nz3 : bits 1) (nz2 : bits 1)))))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))))
- then
- let ui6 : bits 1 := subrange_vec_dec v__460 5 5 in
- let ui53 : bits 3 := subrange_vec_dec v__460 12 10 in
- let ui2 : bits 1 := subrange_vec_dec v__460 6 6 in
- let rs1 : cregidx := subrange_vec_dec v__460 9 7 in
- let rd : cregidx := subrange_vec_dec v__460 4 2 in
- C_LW
- ((concat_vec (ui6 : bits 1) (concat_vec (ui53 : bits 3) (ui2 : bits 1)), rs1, rd))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))))) then
- let ui76 : bits 2 := subrange_vec_dec v__460 6 5 in
- let ui53 : bits 3 := subrange_vec_dec v__460 12 10 in
- let rs1 : cregidx := subrange_vec_dec v__460 9 7 in
- let rd : cregidx := subrange_vec_dec v__460 4 2 in
- C_LD
- ((concat_vec (ui76 : bits 2) (ui53 : bits 3), rs1, rd))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))))
- then
- let ui6 : bits 1 := subrange_vec_dec v__460 5 5 in
- let ui53 : bits 3 := subrange_vec_dec v__460 12 10 in
- let ui2 : bits 1 := subrange_vec_dec v__460 6 6 in
- let rs2 : cregidx := subrange_vec_dec v__460 4 2 in
- let rs1 : cregidx := subrange_vec_dec v__460 9 7 in
- C_SW
- ((concat_vec (ui6 : bits 1) (concat_vec (ui53 : bits 3) (ui2 : bits 1)), rs1, rs2))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))))) then
- let ui76 : bits 2 := subrange_vec_dec v__460 6 5 in
- let ui53 : bits 3 := subrange_vec_dec v__460 12 10 in
- let rs2 : bits 3 := subrange_vec_dec v__460 4 2 in
- let rs1 : bits 3 := subrange_vec_dec v__460 9 7 in
- C_SD
- ((concat_vec (ui76 : bits 2) (ui53 : bits 3), rs1, rs2))
- else if ((andb
- (let rsd : regidx := subrange_vec_dec v__460 11 7 in
- let nzi5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzi40 : bits 5 := subrange_vec_dec v__460 6 2 in
- andb
- (neq_vec (concat_vec nzi5 nzi40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B0;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- let rsd : regidx := subrange_vec_dec v__460 11 7 in
- let nzi5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzi40 : bits 5 := subrange_vec_dec v__460 6 2 in
- C_ADDI
- ((concat_vec (nzi5 : bits 1) (nzi40 : bits 5), rsd))
- else if sumbool_of_bool ((andb (Z.eqb 64 32)
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B0;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))) then
- let i98 : bits 2 := subrange_vec_dec v__460 10 9 in
- let i7 : bits 1 := subrange_vec_dec v__460 6 6 in
- let i6 : bits 1 := subrange_vec_dec v__460 7 7 in
- let i5 : bits 1 := subrange_vec_dec v__460 2 2 in
- let i4 : bits 1 := subrange_vec_dec v__460 11 11 in
- let i31 : bits 3 := subrange_vec_dec v__460 5 3 in
- let i11 : bits 1 := subrange_vec_dec v__460 12 12 in
- let i10 : bits 1 := subrange_vec_dec v__460 8 8 in
- C_JAL
- (concat_vec (i11 : bits 1)
- (concat_vec (i10 : bits 1)
- (concat_vec (i98 : bits 2)
- (concat_vec (i7 : bits 1)
- (concat_vec (i6 : bits 1)
- (concat_vec (i5 : bits 1) (concat_vec (i4 : bits 1) (i31 : bits 3))))))))
- else if sumbool_of_bool ((andb
- (let rsd : regidx := subrange_vec_dec v__460 11 7 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
- (projT1
- (regidx_to_regno zreg)))) (Z.eqb 64 64))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B0;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))) then
- let rsd : regidx := subrange_vec_dec v__460 11 7 in
- let imm5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let imm40 : bits 5 := subrange_vec_dec v__460 6 2 in
- C_ADDIW
- ((concat_vec (imm5 : bits 1) (imm40 : bits 5), rsd))
- else if sumbool_of_bool ((andb
- (let rd : regidx := subrange_vec_dec v__460 11 7 in
- projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))) then
- let rd : regidx := subrange_vec_dec v__460 11 7 in
- let imm5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let imm40 : bits 5 := subrange_vec_dec v__460 6 2 in
- C_LI
- ((concat_vec (imm5 : bits 1) (imm40 : bits 5), rd))
- else if ((andb
- (let nzi9 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzi87 : bits 2 := subrange_vec_dec v__460 4 3 in
- let nzi6 : bits 1 := subrange_vec_dec v__460 5 5 in
- let nzi5 : bits 1 := subrange_vec_dec v__460 2 2 in
- let nzi4 : bits 1 := subrange_vec_dec v__460 6 6 in
- neq_vec (concat_vec nzi9 (concat_vec nzi87 (concat_vec nzi6 (concat_vec nzi5 nzi4))))
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + (2 + (1 + (1 + 1))))))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B1;B1] : mword (15 - 13 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 11 7)
- (vec_of_bits [B0;B0;B0;B1;B0] : mword (11 - 7 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- let nzi9 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzi87 : bits 2 := subrange_vec_dec v__460 4 3 in
- let nzi6 : bits 1 := subrange_vec_dec v__460 5 5 in
- let nzi5 : bits 1 := subrange_vec_dec v__460 2 2 in
- let nzi4 : bits 1 := subrange_vec_dec v__460 6 6 in
- C_ADDI16SP
- (concat_vec (nzi9 : bits 1)
- (concat_vec (nzi87 : bits 2)
- (concat_vec (nzi6 : bits 1) (concat_vec (nzi5 : bits 1) (nzi4 : bits 1)))))
- else if sumbool_of_bool ((andb
- (let rd : regidx := subrange_vec_dec v__460 11 7 in
- let imm17 : bits 1 := subrange_vec_dec v__460 12 12 in
- let imm1612 : bits 5 := subrange_vec_dec v__460 6 2 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno sp))))
- (neq_vec (concat_vec imm17 imm1612)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))) then
- let rd : regidx := subrange_vec_dec v__460 11 7 in
- let imm17 : bits 1 := subrange_vec_dec v__460 12 12 in
- let imm1612 : bits 5 := subrange_vec_dec v__460 6 2 in
- C_LUI
- ((concat_vec (imm17 : bits 1) (imm1612 : bits 5), rd))
- else if ((andb
- (let nzui5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__460 6 2 in
- neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B0;B0] : mword (15 - 13 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 11 10)
- (vec_of_bits [B0;B0] : mword (11 - 10 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let nzui5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__460 6 2 in
- C_SRLI
- ((concat_vec (nzui5 : bits 1) (nzui40 : bits 5), rsd))
- else if ((andb
- (let nzui5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__460 6 2 in
- neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B0;B0] : mword (15 - 13 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 11 10)
- (vec_of_bits [B0;B1] : mword (11 - 10 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let nzui5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__460 6 2 in
- C_SRAI
- ((concat_vec (nzui5 : bits 1) (nzui40 : bits 5), rsd))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B0;B0] : mword (15 - 13 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 11 10)
- (vec_of_bits [B1;B0] : mword (11 - 10 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let i5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let i40 : bits 5 := subrange_vec_dec v__460 6 2 in
- C_ANDI
- ((concat_vec (i5 : bits 1) (i40 : bits 5), rsd))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 10)
- (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 6 5) (vec_of_bits [B0;B0] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let rs2 : cregidx := subrange_vec_dec v__460 4 2 in
- C_SUB
- ((rsd, rs2))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 10)
- (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 6 5) (vec_of_bits [B0;B1] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let rs2 : cregidx := subrange_vec_dec v__460 4 2 in
- C_XOR
- ((rsd, rs2))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 10)
- (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 6 5) (vec_of_bits [B1;B0] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let rs2 : cregidx := subrange_vec_dec v__460 4 2 in
- C_OR
- ((rsd, rs2))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 10)
- (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 6 5) (vec_of_bits [B1;B1] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let rs2 : cregidx := subrange_vec_dec v__460 4 2 in
- C_AND
- ((rsd, rs2))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__460 15 10)
- (vec_of_bits [B1;B0;B0;B1;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 6 5)
- (vec_of_bits [B0;B0] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let rs2 : cregidx := subrange_vec_dec v__460 4 2 in
- C_SUBW
- ((rsd, rs2))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__460 15 10)
- (vec_of_bits [B1;B0;B0;B1;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__460 6 5)
- (vec_of_bits [B0;B1] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- let rsd : cregidx := subrange_vec_dec v__460 9 7 in
- let rs2 : cregidx := subrange_vec_dec v__460 4 2 in
- C_ADDW
- ((rsd, rs2))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B0;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- let i98 : bits 2 := subrange_vec_dec v__460 10 9 in
- let i7 : bits 1 := subrange_vec_dec v__460 6 6 in
- let i6 : bits 1 := subrange_vec_dec v__460 7 7 in
- let i5 : bits 1 := subrange_vec_dec v__460 2 2 in
- let i4 : bits 1 := subrange_vec_dec v__460 11 11 in
- let i31 : bits 3 := subrange_vec_dec v__460 5 3 in
- let i11 : bits 1 := subrange_vec_dec v__460 12 12 in
- let i10 : bits 1 := subrange_vec_dec v__460 8 8 in
- C_J
- (concat_vec (i11 : bits 1)
- (concat_vec (i10 : bits 1)
- (concat_vec (i98 : bits 2)
- (concat_vec (i7 : bits 1)
- (concat_vec (i6 : bits 1)
- (concat_vec (i5 : bits 1) (concat_vec (i4 : bits 1) (i31 : bits 3))))))))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- let rs : cregidx := subrange_vec_dec v__460 9 7 in
- let i8 : bits 1 := subrange_vec_dec v__460 12 12 in
- let i76 : bits 2 := subrange_vec_dec v__460 6 5 in
- let i5 : bits 1 := subrange_vec_dec v__460 2 2 in
- let i43 : bits 2 := subrange_vec_dec v__460 11 10 in
- let i21 : bits 2 := subrange_vec_dec v__460 4 3 in
- C_BEQZ
- ((concat_vec (i8 : bits 1)
- (concat_vec (i76 : bits 2)
- (concat_vec (i5 : bits 1) (concat_vec (i43 : bits 2) (i21 : bits 2)))), rs))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- let rs : cregidx := subrange_vec_dec v__460 9 7 in
- let i8 : bits 1 := subrange_vec_dec v__460 12 12 in
- let i76 : bits 2 := subrange_vec_dec v__460 6 5 in
- let i5 : bits 1 := subrange_vec_dec v__460 2 2 in
- let i43 : bits 2 := subrange_vec_dec v__460 11 10 in
- let i21 : bits 2 := subrange_vec_dec v__460 4 3 in
- C_BNEZ
- ((concat_vec (i8 : bits 1)
- (concat_vec (i76 : bits 2)
- (concat_vec (i5 : bits 1) (concat_vec (i43 : bits 2) (i21 : bits 2)))), rs))
- else if ((andb
- (let rsd : regidx := subrange_vec_dec v__460 11 7 in
- let nzui5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__460 6 2 in
- andb
- (neq_vec (concat_vec nzui5 nzui40)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- (orb (Z.eqb 64 64) (eq_vec nzui5 ((bool_to_bits false) : mword 1))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B0;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1))))))
- then
- let rsd : regidx := subrange_vec_dec v__460 11 7 in
- let nzui5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__460 6 2 in
- C_SLLI
- ((concat_vec (nzui5 : bits 1) (nzui40 : bits 5), rsd))
- else if sumbool_of_bool ((andb
- (let rd : regidx := subrange_vec_dec v__460 11 7 in
- projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- let ui76 : bits 2 := subrange_vec_dec v__460 3 2 in
- let ui5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let ui42 : bits 3 := subrange_vec_dec v__460 6 4 in
- let rd : regidx := subrange_vec_dec v__460 11 7 in
- C_LWSP
- ((concat_vec (ui76 : bits 2) (concat_vec (ui5 : bits 1) (ui42 : bits 3)), rd))
- else if sumbool_of_bool ((andb
- (let rd : regidx := subrange_vec_dec v__460 11 7 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg)))) (Z.eqb 64 64))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B0;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- let ui86 : bits 3 := subrange_vec_dec v__460 4 2 in
- let ui5 : bits 1 := subrange_vec_dec v__460 12 12 in
- let ui43 : bits 2 := subrange_vec_dec v__460 6 5 in
- let rd : regidx := subrange_vec_dec v__460 11 7 in
- C_LDSP
- ((concat_vec (ui86 : bits 3) (concat_vec (ui5 : bits 1) (ui43 : bits 2)), rd))
- else if ((andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))
- then
- let ui76 : bits 2 := subrange_vec_dec v__460 8 7 in
- let ui52 : bits 4 := subrange_vec_dec v__460 12 9 in
- let rs2 : regidx := subrange_vec_dec v__460 6 2 in
- C_SWSP
- ((concat_vec (ui76 : bits 2) (ui52 : bits 4), rs2))
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__460 15 13)
- (vec_of_bits [B1;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- let ui86 : bits 3 := subrange_vec_dec v__460 9 7 in
- let ui53 : bits 3 := subrange_vec_dec v__460 12 10 in
- let rs2 : regidx := subrange_vec_dec v__460 6 2 in
- C_SDSP
- ((concat_vec (ui86 : bits 3) (ui53 : bits 3), rs2))
- else if sumbool_of_bool ((andb
- (let rs1 : regidx := subrange_vec_dec v__460 11 7 in
- projT1
- (neq_int (projT1 (regidx_to_regno rs1))
- (projT1
- (regidx_to_regno zreg))))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 12)
- (vec_of_bits [B1;B0;B0;B0] : mword (15 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__460 6 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B1;B0] : mword (6 - 0 + 1))))))
- then
- let rs1 : regidx := subrange_vec_dec v__460 11 7 in
- C_JR
- (rs1)
- else if sumbool_of_bool ((andb
- (let rs1 : regidx := subrange_vec_dec v__460 11 7 in
- projT1
- (neq_int (projT1 (regidx_to_regno rs1))
+Definition encdec_compressed_backwards (arg_ : mword 16) : M (ast) :=
+ let v__940 := arg_ in
+ (if eq_vec v__940 (Ox"0001" : mword 16) then returnm (C_NOP tt)
+ else if andb
+ (let nz96 : bits 4 := subrange_vec_dec v__940 10 7 in
+ let nz54 : bits 2 := subrange_vec_dec v__940 12 11 in
+ let nz3 : bits 1 := subrange_vec_dec v__940 5 5 in
+ let nz2 : bits 1 := subrange_vec_dec v__940 6 6 in
+ neq_vec (concat_vec nz96 (concat_vec nz54 (concat_vec nz3 nz2)))
+ (Ox"00"
+ : mword (4 + (2 + (1 + 1)))))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"000" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"00" : mword (1 - 0 + 1)))) then
+ let rd : cregidx := subrange_vec_dec v__940 4 2 in
+ let nz96 : bits 4 := subrange_vec_dec v__940 10 7 in
+ let nz54 : bits 2 := subrange_vec_dec v__940 12 11 in
+ let nz3 : bits 1 := subrange_vec_dec v__940 5 5 in
+ let nz2 : bits 1 := subrange_vec_dec v__940 6 6 in
+ returnm (C_ADDI4SPN
+ (rd, concat_vec (nz96 : bits 4)
+ (concat_vec (nz54 : bits 2) (concat_vec (nz3 : bits 1) (nz2 : bits 1)))))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"010" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"00" : mword (1 - 0 + 1))) then
+ let ui6 : bits 1 := subrange_vec_dec v__940 5 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let ui2 : bits 1 := subrange_vec_dec v__940 6 6 in
+ let rs1 : cregidx := subrange_vec_dec v__940 9 7 in
+ let rd : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_LW
+ (concat_vec (ui6 : bits 1) (concat_vec (ui53 : bits 3) (ui2 : bits 1)), rs1, rd))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"00" : mword (1 - 0 + 1))))) then
+ let ui76 : bits 2 := subrange_vec_dec v__940 6 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let rs1 : cregidx := subrange_vec_dec v__940 9 7 in
+ let rd : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_LD (concat_vec (ui76 : bits 2) (ui53 : bits 3), rs1, rd))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"110" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"00" : mword (1 - 0 + 1))) then
+ let ui6 : bits 1 := subrange_vec_dec v__940 5 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let ui2 : bits 1 := subrange_vec_dec v__940 6 6 in
+ let rs2 : cregidx := subrange_vec_dec v__940 4 2 in
+ let rs1 : cregidx := subrange_vec_dec v__940 9 7 in
+ returnm (C_SW
+ (concat_vec (ui6 : bits 1) (concat_vec (ui53 : bits 3) (ui2 : bits 1)), rs1, rs2))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"111" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"00" : mword (1 - 0 + 1))))) then
+ let ui76 : bits 2 := subrange_vec_dec v__940 6 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let rs2 : bits 3 := subrange_vec_dec v__940 4 2 in
+ let rs1 : bits 3 := subrange_vec_dec v__940 9 7 in
+ returnm (C_SD (concat_vec (ui76 : bits 2) (ui53 : bits 3), rs1, rs2))
+ else if andb
+ (let rsd : regidx := subrange_vec_dec v__940 11 7 in
+ let nzi5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzi40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ andb (neq_vec (concat_vec nzi5 nzi40) ('b"000000" : mword (1 + 5)))
+ (neq_vec rsd zreg))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"000" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ let rsd : regidx := subrange_vec_dec v__940 11 7 in
+ let nzi5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzi40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ returnm (C_ADDI (concat_vec (nzi5 : bits 1) (nzi40 : bits 5), rsd))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 32)
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"001" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ let i98 : bits 2 := subrange_vec_dec v__940 10 9 in
+ let i7 : bits 1 := subrange_vec_dec v__940 6 6 in
+ let i6 : bits 1 := subrange_vec_dec v__940 7 7 in
+ let i5 : bits 1 := subrange_vec_dec v__940 2 2 in
+ let i4 : bits 1 := subrange_vec_dec v__940 11 11 in
+ let i31 : bits 3 := subrange_vec_dec v__940 5 3 in
+ let i11 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let i10 : bits 1 := subrange_vec_dec v__940 8 8 in
+ returnm (C_JAL
+ (concat_vec (i11 : bits 1)
+ (concat_vec (i10 : bits 1)
+ (concat_vec (i98 : bits 2)
+ (concat_vec (i7 : bits 1)
+ (concat_vec (i6 : bits 1)
+ (concat_vec (i5 : bits 1)
+ (concat_vec (i4 : bits 1) (i31 : bits 3)))))))))
+ else if sumbool_of_bool
+ (andb
+ (let rsd : regidx := subrange_vec_dec v__940 11 7 in
+ andb (neq_vec rsd zreg) (Z.eqb 64 64))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"001" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ let rsd : regidx := subrange_vec_dec v__940 11 7 in
+ let imm5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let imm40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ returnm (C_ADDIW (concat_vec (imm5 : bits 1) (imm40 : bits 5), rsd))
+ else if andb
+ (let rd : regidx := subrange_vec_dec v__940 11 7 in
+ neq_vec rd zreg)
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"010" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ let rd : regidx := subrange_vec_dec v__940 11 7 in
+ let imm5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let imm40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ returnm (C_LI (concat_vec (imm5 : bits 1) (imm40 : bits 5), rd))
+ else if andb
+ (let nzi9 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzi87 : bits 2 := subrange_vec_dec v__940 4 3 in
+ let nzi6 : bits 1 := subrange_vec_dec v__940 5 5 in
+ let nzi5 : bits 1 := subrange_vec_dec v__940 2 2 in
+ let nzi4 : bits 1 := subrange_vec_dec v__940 6 6 in
+ neq_vec (concat_vec nzi9 (concat_vec nzi87 (concat_vec nzi6 (concat_vec nzi5 nzi4))))
+ ('b"000000"
+ : mword (1 + (2 + (1 + (1 + 1))))))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 11 7) ('b"00010" : mword (11 - 7 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ let nzi9 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzi87 : bits 2 := subrange_vec_dec v__940 4 3 in
+ let nzi6 : bits 1 := subrange_vec_dec v__940 5 5 in
+ let nzi5 : bits 1 := subrange_vec_dec v__940 2 2 in
+ let nzi4 : bits 1 := subrange_vec_dec v__940 6 6 in
+ returnm (C_ADDI16SP
+ (concat_vec (nzi9 : bits 1)
+ (concat_vec (nzi87 : bits 2)
+ (concat_vec (nzi6 : bits 1) (concat_vec (nzi5 : bits 1) (nzi4 : bits 1))))))
+ else if andb
+ (let rd : regidx := subrange_vec_dec v__940 11 7 in
+ let imm17 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let imm1612 : bits 5 := subrange_vec_dec v__940 6 2 in
+ andb (neq_vec rd zreg)
+ (andb (neq_vec rd sp)
+ (neq_vec (concat_vec imm17 imm1612) ('b"000000" : mword (1 + 5)))))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ let rd : regidx := subrange_vec_dec v__940 11 7 in
+ let imm17 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let imm1612 : bits 5 := subrange_vec_dec v__940 6 2 in
+ returnm (C_LUI (concat_vec (imm17 : bits 1) (imm1612 : bits 5), rd))
+ else if andb
+ (let nzui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"100" : mword (15 - 13 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 11 10) ('b"00" : mword (11 - 10 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let nzui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ returnm (C_SRLI (concat_vec (nzui5 : bits 1) (nzui40 : bits 5), rsd))
+ else if andb
+ (let nzui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"100" : mword (15 - 13 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 11 10) ('b"01" : mword (11 - 10 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let nzui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ returnm (C_SRAI (concat_vec (nzui5 : bits 1) (nzui40 : bits 5), rsd))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"100" : mword (15 - 13 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 11 10) ('b"10" : mword (11 - 10 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let i5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let i40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ returnm (C_ANDI (concat_vec (i5 : bits 1) (i40 : bits 5), rsd))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 10) ('b"100011" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 6 5) ('b"00" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let rs2 : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_SUB (rsd, rs2))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 10) ('b"100011" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 6 5) ('b"01" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let rs2 : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_XOR (rsd, rs2))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 10) ('b"100011" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 6 5) ('b"10" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let rs2 : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_OR (rsd, rs2))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 10) ('b"100011" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 6 5) ('b"11" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let rs2 : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_AND (rsd, rs2))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__940 15 10) ('b"100111" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 6 5) ('b"00" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let rs2 : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_SUBW (rsd, rs2))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__940 15 10) ('b"100111" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__940 6 5) ('b"01" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1)))))) then
+ let rsd : cregidx := subrange_vec_dec v__940 9 7 in
+ let rs2 : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_ADDW (rsd, rs2))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"101" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ let i98 : bits 2 := subrange_vec_dec v__940 10 9 in
+ let i7 : bits 1 := subrange_vec_dec v__940 6 6 in
+ let i6 : bits 1 := subrange_vec_dec v__940 7 7 in
+ let i5 : bits 1 := subrange_vec_dec v__940 2 2 in
+ let i4 : bits 1 := subrange_vec_dec v__940 11 11 in
+ let i31 : bits 3 := subrange_vec_dec v__940 5 3 in
+ let i11 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let i10 : bits 1 := subrange_vec_dec v__940 8 8 in
+ returnm (C_J
+ (concat_vec (i11 : bits 1)
+ (concat_vec (i10 : bits 1)
+ (concat_vec (i98 : bits 2)
+ (concat_vec (i7 : bits 1)
+ (concat_vec (i6 : bits 1)
+ (concat_vec (i5 : bits 1)
+ (concat_vec (i4 : bits 1) (i31 : bits 3)))))))))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"110" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ let rs : cregidx := subrange_vec_dec v__940 9 7 in
+ let i8 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let i76 : bits 2 := subrange_vec_dec v__940 6 5 in
+ let i5 : bits 1 := subrange_vec_dec v__940 2 2 in
+ let i43 : bits 2 := subrange_vec_dec v__940 11 10 in
+ let i21 : bits 2 := subrange_vec_dec v__940 4 3 in
+ returnm (C_BEQZ
+ (concat_vec (i8 : bits 1)
+ (concat_vec (i76 : bits 2)
+ (concat_vec (i5 : bits 1) (concat_vec (i43 : bits 2) (i21 : bits 2)))), rs))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"111" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ let rs : cregidx := subrange_vec_dec v__940 9 7 in
+ let i8 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let i76 : bits 2 := subrange_vec_dec v__940 6 5 in
+ let i5 : bits 1 := subrange_vec_dec v__940 2 2 in
+ let i43 : bits 2 := subrange_vec_dec v__940 11 10 in
+ let i21 : bits 2 := subrange_vec_dec v__940 4 3 in
+ returnm (C_BNEZ
+ (concat_vec (i8 : bits 1)
+ (concat_vec (i76 : bits 2)
+ (concat_vec (i5 : bits 1) (concat_vec (i43 : bits 2) (i21 : bits 2)))), rs))
+ else if sumbool_of_bool
+ (andb
+ (let rsd : regidx := subrange_vec_dec v__940 11 7 in
+ let nzui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ andb (neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)))
+ (andb (neq_vec rsd zreg) (orb (Z.eqb 64 64) (eq_vec nzui5 ('b"0" : mword 1)))))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"000" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1))))) then
+ let rsd : regidx := subrange_vec_dec v__940 11 7 in
+ let nzui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__940 6 2 in
+ returnm (C_SLLI (concat_vec (nzui5 : bits 1) (nzui40 : bits 5), rsd))
+ else if andb
+ (let rd : regidx := subrange_vec_dec v__940 11 7 in
+ neq_vec rd zreg)
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"010" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1)))) then
+ let ui76 : bits 2 := subrange_vec_dec v__940 3 2 in
+ let ui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let ui42 : bits 3 := subrange_vec_dec v__940 6 4 in
+ let rd : regidx := subrange_vec_dec v__940 11 7 in
+ returnm (C_LWSP
+ (concat_vec (ui76 : bits 2) (concat_vec (ui5 : bits 1) (ui42 : bits 3)), rd))
+ else if sumbool_of_bool
+ (andb
+ (let rd : regidx := subrange_vec_dec v__940 11 7 in
+ andb (neq_vec rd zreg) (Z.eqb 64 64))
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1))))) then
+ let ui86 : bits 3 := subrange_vec_dec v__940 4 2 in
+ let ui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let ui43 : bits 2 := subrange_vec_dec v__940 6 5 in
+ let rd : regidx := subrange_vec_dec v__940 11 7 in
+ returnm (C_LDSP
+ (concat_vec (ui86 : bits 3) (concat_vec (ui5 : bits 1) (ui43 : bits 2)), rd))
+ else if andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"110" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1))) then
+ let ui76 : bits 2 := subrange_vec_dec v__940 8 7 in
+ let ui52 : bits 4 := subrange_vec_dec v__940 12 9 in
+ let rs2 : regidx := subrange_vec_dec v__940 6 2 in
+ returnm (C_SWSP (concat_vec (ui76 : bits 2) (ui52 : bits 4), rs2))
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"111" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1))))) then
+ let ui86 : bits 3 := subrange_vec_dec v__940 9 7 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let rs2 : regidx := subrange_vec_dec v__940 6 2 in
+ returnm (C_SDSP (concat_vec (ui86 : bits 3) (ui53 : bits 3), rs2))
+ else if andb
+ (let rs1 : regidx := subrange_vec_dec v__940 11 7 in
+ neq_vec rs1 zreg)
+ (andb (eq_vec (subrange_vec_dec v__940 15 12) (Ox"8" : mword (15 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__940 6 0) ('b"0000010" : mword (6 - 0 + 1)))) then
+ let rs1 : regidx := subrange_vec_dec v__940 11 7 in
+ returnm (C_JR rs1)
+ else if andb
+ (let rs1 : regidx := subrange_vec_dec v__940 11 7 in
+ neq_vec rs1 zreg)
+ (andb (eq_vec (subrange_vec_dec v__940 15 12) (Ox"9" : mword (15 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__940 6 0) ('b"0000010" : mword (6 - 0 + 1)))) then
+ let rs1 : regidx := subrange_vec_dec v__940 11 7 in
+ returnm (C_JALR rs1)
+ else if andb
+ (let rs2 : regidx := subrange_vec_dec v__940 6 2 in
+ let rd : regidx := subrange_vec_dec v__940 11 7 in
+ andb (neq_vec rd zreg) (neq_vec rs2 zreg))
+ (andb (eq_vec (subrange_vec_dec v__940 15 12) (Ox"8" : mword (15 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1)))) then
+ let rs2 : regidx := subrange_vec_dec v__940 6 2 in
+ let rd : regidx := subrange_vec_dec v__940 11 7 in
+ returnm (C_MV (rd, rs2))
+ else if eq_vec v__940 (Ox"9002" : mword 16) then returnm (C_EBREAK tt)
+ else if andb
+ (let rsd : regidx := subrange_vec_dec v__940 11 7 in
+ let rs2 : regidx := subrange_vec_dec v__940 6 2 in
+ andb (neq_vec rsd zreg) (neq_vec rs2 zreg))
+ (andb (eq_vec (subrange_vec_dec v__940 15 12) (Ox"9" : mword (15 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1)))) then
+ let rsd : regidx := subrange_vec_dec v__940 11 7 in
+ let rs2 : regidx := subrange_vec_dec v__940 6 2 in
+ returnm (C_ADD (rsd, rs2))
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb (eq_vec (subrange_vec_dec v__940 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__4 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__4 then
+ let ui76 : bits 2 := subrange_vec_dec v__940 3 2 in
+ let ui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let ui42 : bits 3 := subrange_vec_dec v__940 6 4 in
+ let rd : regidx := subrange_vec_dec v__940 11 7 in
+ returnm (C_FLWSP
+ (concat_vec (ui76 : bits 2) (concat_vec (ui5 : bits 1) (ui42 : bits 3)), rd))
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__940 15 13) ('b"111" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"10" : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__9 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__9 then
+ let ui76 : bits 2 := subrange_vec_dec v__940 8 7 in
+ let ui52 : bits 4 := subrange_vec_dec v__940 12 9 in
+ let rs2 : regidx := subrange_vec_dec v__940 6 2 in
+ returnm (C_FSWSP (concat_vec (ui76 : bits 2) (ui52 : bits 4), rs2))
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__940 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"00" : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__14 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__14 then
+ let ui6 : bits 1 := subrange_vec_dec v__940 5 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let ui2 : bits 1 := subrange_vec_dec v__940 6 6 in
+ let rs1 : cregidx := subrange_vec_dec v__940 9 7 in
+ let rd : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_FLW
+ (concat_vec (ui6 : bits 1) (concat_vec (ui53 : bits 3) (ui2 : bits 1)), rs1, rd))
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__940 15 13)
+ ('b"111"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0) ('b"00" : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__19 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__19 then
+ let ui6 : bits 1 := subrange_vec_dec v__940 5 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let ui2 : bits 1 := subrange_vec_dec v__940 6 6 in
+ let rs2 : cregidx := subrange_vec_dec v__940 4 2 in
+ let rs1 : cregidx := subrange_vec_dec v__940 9 7 in
+ returnm (C_FSW
+ (concat_vec (ui6 : bits 1)
+ (concat_vec (ui53 : bits 3) (ui2 : bits 1)), rs1, rs2))
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__940 15 13)
+ ('b"001"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0)
+ ('b"10"
+ : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__24 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__24 then
+ let ui86 : bits 3 := subrange_vec_dec v__940 4 2 in
+ let ui5 : bits 1 := subrange_vec_dec v__940 12 12 in
+ let ui43 : bits 2 := subrange_vec_dec v__940 6 5 in
+ let rd : regidx := subrange_vec_dec v__940 11 7 in
+ returnm (C_FLDSP
+ (concat_vec (ui86 : bits 3)
+ (concat_vec (ui5 : bits 1) (ui43 : bits 2)), rd))
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__940 15 13)
+ ('b"101"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0)
+ ('b"10"
+ : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__29 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__29 then
+ let ui86 : bits 3 := subrange_vec_dec v__940 9 7 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let rs2 : regidx := subrange_vec_dec v__940 6 2 in
+ returnm (C_FSDSP (concat_vec (ui86 : bits 3) (ui53 : bits 3), rs2))
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
(projT1
- (regidx_to_regno zreg))))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 12)
- (vec_of_bits [B1;B0;B0;B1] : mword (15 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__460 6 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B1;B0] : mword (6 - 0 + 1))))))
- then
- let rs1 : regidx := subrange_vec_dec v__460 11 7 in
- C_JALR
- (rs1)
- else if sumbool_of_bool ((andb
- (let rs2 : regidx := subrange_vec_dec v__460 6 2 in
- let rd : regidx := subrange_vec_dec v__460 11 7 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 12)
- (vec_of_bits [B1;B0;B0;B0] : mword (15 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- let rs2 : regidx := subrange_vec_dec v__460 6 2 in
- let rd : regidx := subrange_vec_dec v__460 11 7 in
- C_MV
- ((rd, rs2))
- else if ((eq_vec v__460
- (vec_of_bits [B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 16))) then
- C_EBREAK
- (tt)
- else if sumbool_of_bool ((andb
- (let rsd : regidx := subrange_vec_dec v__460 11 7 in
- let rs2 : regidx := subrange_vec_dec v__460 6 2 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__940 15 13)
+ ('b"001"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__34 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__34 then
+ let ui76 : bits 2 := subrange_vec_dec v__940 6 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let rs1 : cregidx := subrange_vec_dec v__940 9 7 in
+ let rd : cregidx := subrange_vec_dec v__940 4 2 in
+ returnm (C_FLD (concat_vec (ui76 : bits 2) (ui53 : bits 3), rs1, rd))
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
(projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__460 15 12)
- (vec_of_bits [B1;B0;B0;B1] : mword (15 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__460 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- let rsd : regidx := subrange_vec_dec v__460 11 7 in
- let rs2 : regidx := subrange_vec_dec v__460 6 2 in
- C_ADD
- ((rsd, rs2))
- else C_ILLEGAL (v__460).
-
-Definition encdec_compressed_forwards_matches (arg_ : ast)
-: bool :=
-
- match arg_ with
- | C_NOP (tt) => true
- | C_ADDI4SPN ((rd, v__596)) =>
- if ((let nz96 : bits 4 := subrange_vec_dec v__596 7 4 in
- let nz96 : bits 4 := subrange_vec_dec v__596 7 4 in
- let nz54 : bits 2 := subrange_vec_dec v__596 3 2 in
- let nz3 : bits 1 := subrange_vec_dec v__596 1 1 in
- let nz2 : bits 1 := subrange_vec_dec v__596 0 0 in
- neq_vec (concat_vec nz96 (concat_vec nz54 (concat_vec nz3 nz2)))
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword (4 + (2 + (1 + 1)))))) then
- true
- else false
- | C_LW ((v__597, rs1, rd)) => true
- | C_LD ((v__598, rs1, rd)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_SW ((v__599, rs1, rs2)) => true
- | C_SD ((v__600, rs1, rs2)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_ADDI ((v__601, rsd)) =>
- if ((let nzi5 : bits 1 := subrange_vec_dec v__601 5 5 in
- let nzi5 : bits 1 := subrange_vec_dec v__601 5 5 in
- let nzi40 : bits 5 := subrange_vec_dec v__601 4 0 in
- andb (neq_vec (concat_vec nzi5 nzi40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- : bool))) then
- true
- else false
- | C_JAL (v__602) => if sumbool_of_bool ((Z.eqb 64 32)) then true else false
- | C_ADDIW ((v__603, rsd)) =>
- if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
- (projT1
- (regidx_to_regno zreg)))) (Z.eqb 64 64))) then
- true
- else false
- | C_LI ((v__604, rd)) =>
- if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg)))))
- then
- true
- else false
- | C_ADDI16SP (v__605) =>
- if ((let nzi9 : bits 1 := subrange_vec_dec v__605 5 5 in
- let nzi9 : bits 1 := subrange_vec_dec v__605 5 5 in
- let nzi87 : bits 2 := subrange_vec_dec v__605 4 3 in
- let nzi6 : bits 1 := subrange_vec_dec v__605 2 2 in
- let nzi5 : bits 1 := subrange_vec_dec v__605 1 1 in
- let nzi4 : bits 1 := subrange_vec_dec v__605 0 0 in
- neq_vec (concat_vec nzi9 (concat_vec nzi87 (concat_vec nzi6 (concat_vec nzi5 nzi4))))
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + (2 + (1 + (1 + 1))))))) then
- true
- else false
- | C_LUI ((v__606, rd)) =>
- if sumbool_of_bool ((let imm17 : bits 1 := subrange_vec_dec v__606 5 5 in
- let imm17 : bits 1 := subrange_vec_dec v__606 5 5 in
- let imm1612 : bits 5 := subrange_vec_dec v__606 4 0 in
- andb (projT1 (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg))))
- ((andb (projT1 (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno sp))))
- (neq_vec (concat_vec imm17 imm1612) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5))))
- : bool))) then
- true
- else false
- | C_SRLI ((v__607, rsd)) =>
- if ((let nzui5 : bits 1 := subrange_vec_dec v__607 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__607 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__607 4 0 in
- neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))) then
- true
- else false
- | C_SRAI ((v__608, rsd)) =>
- if ((let nzui5 : bits 1 := subrange_vec_dec v__608 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__608 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__608 4 0 in
- neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))) then
- true
- else false
- | C_ANDI ((v__609, rsd)) => true
- | C_SUB ((rsd, rs2)) => true
- | C_XOR ((rsd, rs2)) => true
- | C_OR ((rsd, rs2)) => true
- | C_AND ((rsd, rs2)) => true
- | C_SUBW ((rsd, rs2)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_ADDW ((rsd, rs2)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_J (v__610) => true
- | C_BEQZ ((v__611, rs)) => true
- | C_BNEZ ((v__612, rs)) => true
- | C_SLLI ((v__613, rsd)) =>
- if ((let nzui5 : bits 1 := subrange_vec_dec v__613 5 5 in
- let nzui5 : bits 1 := subrange_vec_dec v__613 5 5 in
- let nzui40 : bits 5 := subrange_vec_dec v__613 4 0 in
- andb (neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- ((andb (projT1 (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- (orb (Z.eqb 64 64) (eq_vec nzui5 ((bool_to_bits false) : mword 1))))
- : bool))) then
- true
- else false
- | C_LWSP ((v__614, rd)) =>
- if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg)))))
- then
- true
- else false
- | C_LDSP ((v__615, rd)) =>
- if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg))))
- (Z.eqb 64 64))) then
- true
- else false
- | C_SWSP ((v__616, rs2)) => true
- | C_SDSP ((v__617, rs2)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_JR (rs1) =>
- if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rs1)) (projT1 (regidx_to_regno zreg)))))
- then
- true
- else false
- | C_JALR (rs1) =>
- if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rs1)) (projT1 (regidx_to_regno zreg)))))
- then
- true
- else false
- | C_MV ((rd, rs2)) =>
- if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))) then
- true
- else false
- | C_EBREAK (tt) => true
- | C_ADD ((rsd, rs2)) =>
- if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
- (projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))) then
- true
- else false
- | C_ILLEGAL (s) => true
- | _ => false
- end.
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) ||
+ (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool))
+ ((haveDExt tt)
+ : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__940 15 13)
+ ('b"101"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__940 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__39 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__39 then
+ let ui76 : bits 2 := subrange_vec_dec v__940 6 5 in
+ let ui53 : bits 3 := subrange_vec_dec v__940 12 10 in
+ let rs2 : bits 3 := subrange_vec_dec v__940 4 2 in
+ let rs1 : bits 3 := subrange_vec_dec v__940 9 7 in
+ C_FSD (concat_vec (ui76 : bits 2) (ui53 : bits 3), rs1, rs2)
+ else C_ILLEGAL v__940))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast).
-Definition encdec_compressed_backwards_matches (arg_ : mword 16)
-: bool :=
-
- let v__618 := arg_ in
- if ((eq_vec v__618 (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1] : mword 16)))
- then
- true
- else if ((andb
- (let nz96 : bits 4 := subrange_vec_dec v__618 10 7 in
- let nz54 : bits 2 := subrange_vec_dec v__618 12 11 in
- let nz3 : bits 1 := subrange_vec_dec v__618 5 5 in
- let nz2 : bits 1 := subrange_vec_dec v__618 6 6 in
- neq_vec (concat_vec nz96 (concat_vec nz54 (concat_vec nz3 nz2)))
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword (4 + (2 + (1 + 1)))))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B0;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B0] : mword (1 - 0 + 1))))))
- then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))))
- then
- true
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))))) then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))))
- then
- true
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))))) then
- true
- else if ((andb
- (let rsd : regidx := subrange_vec_dec v__618 11 7 in
- let nzi5 : bits 1 := subrange_vec_dec v__618 12 12 in
- let nzi40 : bits 5 := subrange_vec_dec v__618 6 2 in
- andb
- (neq_vec (concat_vec nzi5 nzi40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B0;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- true
- else if sumbool_of_bool ((andb (Z.eqb 64 32)
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B0;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))) then
- true
- else if sumbool_of_bool ((andb
- (let rsd : regidx := subrange_vec_dec v__618 11 7 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
- (projT1
- (regidx_to_regno zreg)))) (Z.eqb 64 64))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B0;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))) then
- true
- else if sumbool_of_bool ((andb
- (let rd : regidx := subrange_vec_dec v__618 11 7 in
- projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))) then
- true
- else if ((andb
- (let nzi9 : bits 1 := subrange_vec_dec v__618 12 12 in
- let nzi87 : bits 2 := subrange_vec_dec v__618 4 3 in
- let nzi6 : bits 1 := subrange_vec_dec v__618 5 5 in
- let nzi5 : bits 1 := subrange_vec_dec v__618 2 2 in
- let nzi4 : bits 1 := subrange_vec_dec v__618 6 6 in
- neq_vec (concat_vec nzi9 (concat_vec nzi87 (concat_vec nzi6 (concat_vec nzi5 nzi4))))
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + (2 + (1 + (1 + 1))))))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B1;B1] : mword (15 - 13 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 11 7)
- (vec_of_bits [B0;B0;B0;B1;B0] : mword (11 - 7 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- true
- else if sumbool_of_bool ((andb
- (let rd : regidx := subrange_vec_dec v__618 11 7 in
- let imm17 : bits 1 := subrange_vec_dec v__618 12 12 in
- let imm1612 : bits 5 := subrange_vec_dec v__618 6 2 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno sp))))
- (neq_vec (concat_vec imm17 imm1612)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))) then
- true
- else if ((andb
- (let nzui5 : bits 1 := subrange_vec_dec v__618 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__618 6 2 in
- neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B0;B0] : mword (15 - 13 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 11 10)
- (vec_of_bits [B0;B0] : mword (11 - 10 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- true
- else if ((andb
- (let nzui5 : bits 1 := subrange_vec_dec v__618 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__618 6 2 in
- neq_vec (concat_vec nzui5 nzui40) (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B0;B0] : mword (15 - 13 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 11 10)
- (vec_of_bits [B0;B1] : mword (11 - 10 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B0;B0] : mword (15 - 13 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 11 10)
- (vec_of_bits [B1;B0] : mword (11 - 10 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 10)
- (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 6 5) (vec_of_bits [B0;B0] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 10)
- (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 6 5) (vec_of_bits [B0;B1] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 10)
- (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 6 5) (vec_of_bits [B1;B0] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 10)
- (vec_of_bits [B1;B0;B0;B0;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 6 5) (vec_of_bits [B1;B1] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))
- then
- true
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__618 15 10)
- (vec_of_bits [B1;B0;B0;B1;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 6 5)
- (vec_of_bits [B0;B0] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- true
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__618 15 10)
- (vec_of_bits [B1;B0;B0;B1;B1;B1] : mword (15 - 10 + 1)))
- (andb
- (eq_vec (subrange_vec_dec v__618 6 5)
- (vec_of_bits [B0;B1] : mword (6 - 5 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))))) then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B0;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- true
- else if ((andb
- (let rsd : regidx := subrange_vec_dec v__618 11 7 in
- let nzui5 : bits 1 := subrange_vec_dec v__618 12 12 in
- let nzui40 : bits 5 := subrange_vec_dec v__618 6 2 in
- andb
- (neq_vec (concat_vec nzui5 nzui40)
- (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword (1 + 5)))
- ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- (orb (Z.eqb 64 64) (eq_vec nzui5 ((bool_to_bits false) : mword 1))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B0;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1))))))
- then
- true
- else if sumbool_of_bool ((andb
- (let rd : regidx := subrange_vec_dec v__618 11 7 in
- projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- true
- else if sumbool_of_bool ((andb
- (let rd : regidx := subrange_vec_dec v__618 11 7 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg)))) (Z.eqb 64 64))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B0;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- true
- else if ((andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B1;B0] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))
- then
- true
- else if sumbool_of_bool ((andb (Z.eqb 64 64)
- (andb
- (eq_vec (subrange_vec_dec v__618 15 13)
- (vec_of_bits [B1;B1;B1] : mword (15 - 13 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- true
- else if sumbool_of_bool ((andb
- (let rs1 : regidx := subrange_vec_dec v__618 11 7 in
- projT1
- (neq_int (projT1 (regidx_to_regno rs1))
- (projT1
- (regidx_to_regno zreg))))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 12)
- (vec_of_bits [B1;B0;B0;B0] : mword (15 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__618 6 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B1;B0] : mword (6 - 0 + 1))))))
- then
- true
- else if sumbool_of_bool ((andb
- (let rs1 : regidx := subrange_vec_dec v__618 11 7 in
- projT1
- (neq_int (projT1 (regidx_to_regno rs1))
+Definition encdec_compressed_forwards_matches (arg_ : ast) : M (bool) :=
+ (match arg_ with
+ | C_NOP tt => returnm true
+ | C_ADDI4SPN (rd, v__1100) =>
+ returnm (if let nz96 : bits 4 := subrange_vec_dec v__1100 7 4 in
+ let nz96 : bits 4 := subrange_vec_dec v__1100 7 4 in
+ let nz54 : bits 2 := subrange_vec_dec v__1100 3 2 in
+ let nz3 : bits 1 := subrange_vec_dec v__1100 1 1 in
+ let nz2 : bits 1 := subrange_vec_dec v__1100 0 0 in
+ neq_vec (concat_vec nz96 (concat_vec nz54 (concat_vec nz3 nz2)))
+ (Ox"00"
+ : mword (4 + (2 + (1 + 1)))) then
+ true
+ else false)
+ | C_LW (v__1101, rs1, rd) => returnm true
+ | C_LD (v__1102, rs1, rd) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | C_SW (v__1103, rs1, rs2) => returnm true
+ | C_SD (v__1104, rs1, rs2) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | C_ADDI (v__1105, rsd) =>
+ returnm (if let nzi5 : bits 1 := subrange_vec_dec v__1105 5 5 in
+ let nzi5 : bits 1 := subrange_vec_dec v__1105 5 5 in
+ let nzi40 : bits 5 := subrange_vec_dec v__1105 4 0 in
+ andb (neq_vec (concat_vec nzi5 nzi40) ('b"000000" : mword (1 + 5)))
+ (neq_vec rsd zreg) then
+ true
+ else false)
+ | C_JAL v__1106 => returnm (if sumbool_of_bool (Z.eqb 64 32) then true else false)
+ | C_ADDIW (v__1107, rsd) =>
+ returnm (if sumbool_of_bool (andb (neq_vec rsd zreg) (Z.eqb 64 64)) then true else false)
+ | C_LI (v__1108, rd) => returnm (if neq_vec rd zreg then true else false)
+ | C_ADDI16SP v__1109 =>
+ returnm (if let nzi9 : bits 1 := subrange_vec_dec v__1109 5 5 in
+ let nzi9 : bits 1 := subrange_vec_dec v__1109 5 5 in
+ let nzi87 : bits 2 := subrange_vec_dec v__1109 4 3 in
+ let nzi6 : bits 1 := subrange_vec_dec v__1109 2 2 in
+ let nzi5 : bits 1 := subrange_vec_dec v__1109 1 1 in
+ let nzi4 : bits 1 := subrange_vec_dec v__1109 0 0 in
+ neq_vec
+ (concat_vec nzi9 (concat_vec nzi87 (concat_vec nzi6 (concat_vec nzi5 nzi4))))
+ ('b"000000"
+ : mword (1 + (2 + (1 + (1 + 1))))) then
+ true
+ else false)
+ | C_LUI (v__1110, rd) =>
+ returnm (if let imm17 : bits 1 := subrange_vec_dec v__1110 5 5 in
+ let imm17 : bits 1 := subrange_vec_dec v__1110 5 5 in
+ let imm1612 : bits 5 := subrange_vec_dec v__1110 4 0 in
+ andb (neq_vec rd zreg)
+ (andb (neq_vec rd sp)
+ (neq_vec (concat_vec imm17 imm1612) ('b"000000" : mword (1 + 5)))) then
+ true
+ else false)
+ | C_SRLI (v__1111, rsd) =>
+ returnm (if let nzui5 : bits 1 := subrange_vec_dec v__1111 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__1111 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__1111 4 0 in
+ neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)) then
+ true
+ else false)
+ | C_SRAI (v__1112, rsd) =>
+ returnm (if let nzui5 : bits 1 := subrange_vec_dec v__1112 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__1112 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__1112 4 0 in
+ neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)) then
+ true
+ else false)
+ | C_ANDI (v__1113, rsd) => returnm true
+ | C_SUB (rsd, rs2) => returnm true
+ | C_XOR (rsd, rs2) => returnm true
+ | C_OR (rsd, rs2) => returnm true
+ | C_AND (rsd, rs2) => returnm true
+ | C_SUBW (rsd, rs2) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | C_ADDW (rsd, rs2) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | C_J v__1114 => returnm true
+ | C_BEQZ (v__1115, rs) => returnm true
+ | C_BNEZ (v__1116, rs) => returnm true
+ | C_SLLI (v__1117, rsd) =>
+ returnm (if sumbool_of_bool
+ (let nzui5 : bits 1 := subrange_vec_dec v__1117 5 5 in
+ let nzui5 : bits 1 := subrange_vec_dec v__1117 5 5 in
+ let nzui40 : bits 5 := subrange_vec_dec v__1117 4 0 in
+ andb (neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)))
+ (andb (neq_vec rsd zreg) (orb (Z.eqb 64 64) (eq_vec nzui5 ('b"0" : mword 1)))))
+ then
+ true
+ else false)
+ | C_LWSP (v__1118, rd) => returnm (if neq_vec rd zreg then true else false)
+ | C_LDSP (v__1119, rd) =>
+ returnm (if sumbool_of_bool (andb (neq_vec rd zreg) (Z.eqb 64 64)) then true else false)
+ | C_SWSP (v__1120, rs2) => returnm true
+ | C_SDSP (v__1121, rs2) => returnm (if sumbool_of_bool (Z.eqb 64 64) then true else false)
+ | C_JR rs1 => returnm (if neq_vec rs1 zreg then true else false)
+ | C_JALR rs1 => returnm (if neq_vec rs1 zreg then true else false)
+ | C_MV (rd, rs2) => returnm (if andb (neq_vec rd zreg) (neq_vec rs2 zreg) then true else false)
+ | C_EBREAK tt => returnm true
+ | C_ADD (rsd, rs2) =>
+ returnm (if andb (neq_vec rsd zreg) (neq_vec rs2 zreg) then true else false)
+ | C_FLWSP (v__1122, rd) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__3 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__3 then true else false)
+ | C_FSWSP (v__1123, rs2) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__7 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__7 then true else false)
+ | C_FLW (v__1124, rs1, rd) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__11 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__11 then true else false)
+ | C_FSW (v__1125, rs1, rs2) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__15 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__15 then true else false)
+ | C_FLDSP (v__1126, rd) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__19 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__19 then true else false)
+ | C_FSDSP (v__1127, rs2) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__23 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__23 then true else false)
+ | C_FLD (v__1128, rs1, rd) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__27 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__27 then true else false)
+ | C_FSD (v__1129, rs1, rs2) =>
+ (and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__31 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__31 then true else false)
+ | C_ILLEGAL s => returnm true
+ | _ => returnm false
+ end)
+ : M (bool).
+
+Definition encdec_compressed_backwards_matches (arg_ : mword 16) : M (bool) :=
+ let v__1130 := arg_ in
+ (if eq_vec v__1130 (Ox"0001" : mword 16) then returnm true
+ else if andb
+ (let nz96 : bits 4 := subrange_vec_dec v__1130 10 7 in
+ let nz54 : bits 2 := subrange_vec_dec v__1130 12 11 in
+ let nz3 : bits 1 := subrange_vec_dec v__1130 5 5 in
+ let nz2 : bits 1 := subrange_vec_dec v__1130 6 6 in
+ neq_vec (concat_vec nz96 (concat_vec nz54 (concat_vec nz3 nz2)))
+ (Ox"00"
+ : mword (4 + (2 + (1 + 1)))))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"000" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"00" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"010" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"00" : mword (1 - 0 + 1))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"00" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"110" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"00" : mword (1 - 0 + 1))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"111" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"00" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb
+ (let rsd : regidx := subrange_vec_dec v__1130 11 7 in
+ let nzi5 : bits 1 := subrange_vec_dec v__1130 12 12 in
+ let nzi40 : bits 5 := subrange_vec_dec v__1130 6 2 in
+ andb (neq_vec (concat_vec nzi5 nzi40) ('b"000000" : mword (1 + 5)))
+ (neq_vec rsd zreg))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"000" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 32)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"001" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb
+ (let rsd : regidx := subrange_vec_dec v__1130 11 7 in
+ andb (neq_vec rsd zreg) (Z.eqb 64 64))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"001" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb
+ (let rd : regidx := subrange_vec_dec v__1130 11 7 in
+ neq_vec rd zreg)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"010" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if andb
+ (let nzi9 : bits 1 := subrange_vec_dec v__1130 12 12 in
+ let nzi87 : bits 2 := subrange_vec_dec v__1130 4 3 in
+ let nzi6 : bits 1 := subrange_vec_dec v__1130 5 5 in
+ let nzi5 : bits 1 := subrange_vec_dec v__1130 2 2 in
+ let nzi4 : bits 1 := subrange_vec_dec v__1130 6 6 in
+ neq_vec (concat_vec nzi9 (concat_vec nzi87 (concat_vec nzi6 (concat_vec nzi5 nzi4))))
+ ('b"000000"
+ : mword (1 + (2 + (1 + (1 + 1))))))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 11 7) ('b"00010" : mword (11 - 7 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb
+ (let rd : regidx := subrange_vec_dec v__1130 11 7 in
+ let imm17 : bits 1 := subrange_vec_dec v__1130 12 12 in
+ let imm1612 : bits 5 := subrange_vec_dec v__1130 6 2 in
+ andb (neq_vec rd zreg)
+ (andb (neq_vec rd sp)
+ (neq_vec (concat_vec imm17 imm1612) ('b"000000" : mword (1 + 5)))))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if andb
+ (let nzui5 : bits 1 := subrange_vec_dec v__1130 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__1130 6 2 in
+ neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"100" : mword (15 - 13 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 11 10) ('b"00" : mword (11 - 10 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb
+ (let nzui5 : bits 1 := subrange_vec_dec v__1130 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__1130 6 2 in
+ neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"100" : mword (15 - 13 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 11 10) ('b"01" : mword (11 - 10 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"100" : mword (15 - 13 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 11 10) ('b"10" : mword (11 - 10 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 10) ('b"100011" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 6 5) ('b"00" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 10) ('b"100011" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 6 5) ('b"01" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 10) ('b"100011" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 6 5) ('b"10" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 10) ('b"100011" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 6 5) ('b"11" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 10) ('b"100111" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 6 5) ('b"00" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 10) ('b"100111" : mword (15 - 10 + 1)))
+ (andb (eq_vec (subrange_vec_dec v__1130 6 5) ('b"01" : mword (6 - 5 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1)))))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"101" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"110" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"111" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb
+ (let rsd : regidx := subrange_vec_dec v__1130 11 7 in
+ let nzui5 : bits 1 := subrange_vec_dec v__1130 12 12 in
+ let nzui40 : bits 5 := subrange_vec_dec v__1130 6 2 in
+ andb (neq_vec (concat_vec nzui5 nzui40) ('b"000000" : mword (1 + 5)))
+ (andb (neq_vec rsd zreg) (orb (Z.eqb 64 64) (eq_vec nzui5 ('b"0" : mword 1)))))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"000" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb
+ (let rd : regidx := subrange_vec_dec v__1130 11 7 in
+ neq_vec rd zreg)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"010" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb
+ (let rd : regidx := subrange_vec_dec v__1130 11 7 in
+ andb (neq_vec rd zreg) (Z.eqb 64 64))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"110" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1))) then
+ returnm true
+ else if sumbool_of_bool
+ (andb (Z.eqb 64 64)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"111" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1))))) then
+ returnm true
+ else if andb
+ (let rs1 : regidx := subrange_vec_dec v__1130 11 7 in
+ neq_vec rs1 zreg)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 12) (Ox"8" : mword (15 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 6 0) ('b"0000010" : mword (6 - 0 + 1)))) then
+ returnm true
+ else if andb
+ (let rs1 : regidx := subrange_vec_dec v__1130 11 7 in
+ neq_vec rs1 zreg)
+ (andb (eq_vec (subrange_vec_dec v__1130 15 12) (Ox"9" : mword (15 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 6 0) ('b"0000010" : mword (6 - 0 + 1)))) then
+ returnm true
+ else if andb
+ (let rs2 : regidx := subrange_vec_dec v__1130 6 2 in
+ let rd : regidx := subrange_vec_dec v__1130 11 7 in
+ andb (neq_vec rd zreg) (neq_vec rs2 zreg))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 12) (Ox"8" : mword (15 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1)))) then
+ returnm true
+ else if eq_vec v__1130 (Ox"9002" : mword 16) then returnm true
+ else if andb
+ (let rsd : regidx := subrange_vec_dec v__1130 11 7 in
+ let rs2 : regidx := subrange_vec_dec v__1130 6 2 in
+ andb (neq_vec rsd zreg) (neq_vec rs2 zreg))
+ (andb (eq_vec (subrange_vec_dec v__1130 15 12) (Ox"9" : mword (15 - 12 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1)))) then
+ returnm true
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb (eq_vec (subrange_vec_dec v__1130 15 13) ('b"011" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__4 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__4 then returnm true
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__1130 15 13) ('b"111" : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"10" : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__9 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__9 then returnm true
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__1130 15 13)
+ ('b"011"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"00" : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__14 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__14 then returnm true
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (Z.eqb 64 32)
+ : {_bool : bool & ArithFact (Bool.eqb (64 =? 32) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb (64 =?
+ 32) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveFExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 && (64 =? 32)) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__1130 15 13)
+ ('b"111"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0) ('b"00" : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)})) >>= fun '(existT _ w__19 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ (64 =? 32)) _bool = true)}) =>
+ (if sumbool_of_bool w__19 then returnm true
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__1130 15 13)
+ ('b"001"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0)
+ ('b"10"
+ : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__24 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__24 then returnm true
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
+ (projT1
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__1130 15 13)
+ ('b"101"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0)
+ ('b"10"
+ : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__29 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__29 then returnm true
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
(projT1
- (regidx_to_regno zreg))))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 12)
- (vec_of_bits [B1;B0;B0;B1] : mword (15 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__618 6 0)
- (vec_of_bits [B0;B0;B0;B0;B0;B1;B0] : mword (6 - 0 + 1))))))
- then
- true
- else if sumbool_of_bool ((andb
- (let rs2 : regidx := subrange_vec_dec v__618 6 2 in
- let rd : regidx := subrange_vec_dec v__618 11 7 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 12)
- (vec_of_bits [B1;B0;B0;B0] : mword (15 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- true
- else if ((eq_vec v__618
- (vec_of_bits [B1;B0;B0;B1;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B1;B0] : mword 16))) then
- true
- else if sumbool_of_bool ((andb
- (let rsd : regidx := subrange_vec_dec v__618 11 7 in
- let rs2 : regidx := subrange_vec_dec v__618 6 2 in
- andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) || (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool)) ((haveDExt tt) : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__1130 15 13)
+ ('b"001"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__34 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ (if sumbool_of_bool w__34 then returnm true
+ else
+ (and_boolMP
+ (((and_boolMP
+ ((returnm (build_ex
(projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))
- (andb
- (eq_vec (subrange_vec_dec v__618 15 12)
- (vec_of_bits [B1;B0;B0;B1] : mword (15 - 12 + 1)))
- (eq_vec (subrange_vec_dec v__618 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))) then
- true
- else true.
+ (build_ex
+ (orb (Z.eqb 64 32) (Z.eqb 64 64))
+ : {_bool : bool & ArithFact (Bool.eqb ((64 =? 32) ||
+ (64 =? 64)) _bool)})))) : M ({_bool : bool & ArithFact (Bool.eqb ((64 =?
+ 32) ||
+ (64 =? 64)) _bool)}))
+ (build_trivial_ex
+ ((and_boolM ((haveRVC tt) : M (bool))
+ ((haveDExt tt)
+ : M (bool)))
+ : M (bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}))
+ (build_trivial_ex
+ (returnm ((andb
+ (eq_vec (subrange_vec_dec v__1130 15 13)
+ ('b"101"
+ : mword (15 - 13 + 1)))
+ (eq_vec (subrange_vec_dec v__1130 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1))))
+ : bool))) : M ({_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)})) >>= fun '(existT _ w__39 _ : {_bool : bool & ArithFactP (exists simp_0 , Bool.eqb (simp_0 &&
+ ((64 =? 32) || (64 =? 64))) _bool = true)}) =>
+ returnm (if sumbool_of_bool w__39 then true else true))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool).
-Definition execute_WFI '(tt : unit)
-: M (Retired) :=
-
+Definition execute_WFI '(tt : unit) : M (Retired) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
(match w__0 with
- | Machine => (platform_wfi tt) >> returnm (RETIRE_SUCCESS : Retired)
+ | Machine => (platform_wfi tt) >> returnm RETIRE_SUCCESS
| Supervisor =>
read_reg mstatus_ref >>= fun w__1 : Mstatus =>
- (if ((eq_vec (_get_Mstatus_TW w__1) ((bool_to_bits true) : mword 1))) then
- (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired)
- else (platform_wfi tt) >> returnm (RETIRE_SUCCESS : Retired))
+ (if eq_vec (_get_Mstatus_TW w__1) ('b"1" : mword 1) then
+ (handle_illegal tt) >> returnm RETIRE_FAIL
+ else (platform_wfi tt) >> returnm RETIRE_SUCCESS)
: M (Retired)
- | User => (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired)
+ | User => (handle_illegal tt) >> returnm RETIRE_FAIL
end)
: M (Retired).
-Definition execute_UTYPE (imm : mword 20) (rd : mword 5) (op : uop)
-: M (Retired) :=
-
- let off : xlenbits :=
- EXTS 64 (concat_vec imm (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0;B0] : mword 12)) in
+Definition execute_UTYPE (imm : mword 20) (rd : mword 5) (op : uop) : M (Retired) :=
+ let off : xlenbits := EXTS 64 (concat_vec imm (Ox"000" : mword 12)) in
(match op with
- | RISCV_LUI => returnm (off : mword 64)
- | RISCV_AUIPC =>
- (get_arch_pc tt) >>= fun w__0 : mword 64 => returnm ((add_vec w__0 off) : mword 64)
+ | RISCV_LUI => returnm off
+ | RISCV_AUIPC => (get_arch_pc tt) >>= fun w__0 : mword 64 => returnm (add_vec w__0 off)
end) >>= fun ret : xlenbits =>
- (wX (projT1 (regidx_to_regno rd)) ret) >> returnm (RETIRE_SUCCESS : Retired).
+ (wX_bits rd ret) >> returnm RETIRE_SUCCESS.
-Definition execute_URET '(tt : unit)
-: M (Retired) :=
-
+Definition execute_URET '(tt : unit) : M (Retired) :=
(haveUsrMode tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool ((negb w__0)) then (handle_illegal tt) : M (unit)
+ (if sumbool_of_bool (negb w__0) then (handle_illegal tt) : M (unit)
+ else if negb (ext_check_xret_priv User) then returnm (ext_fail_xret_priv tt)
else
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
((read_reg PC_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
- (exception_handler w__1 (CTL_URET (tt)) w__2) >>= fun w__3 : mword 64 =>
+ (exception_handler w__1 (CTL_URET tt) w__2) >>= fun w__3 : mword 64 =>
(set_next_pc w__3)
: M (unit)) >>
- returnm (RETIRE_FAIL
- : Retired).
+ returnm RETIRE_FAIL.
+
+Definition execute_STORE_FP (imm : mword 12) (rs2 : mword 5) (rs1 : mword 5) (width : word_width)
+: M (Retired) :=
+ let offset : xlenbits := EXTS 64 imm in
+ let '(aq, rl, con) := (false, false, false) in
+ (ext_data_get_addr rs1 offset (Write Data) width) >>= fun w__0 : Ext_DataAddr_Check unit =>
+ (match w__0 with
+ | Ext_DataAddr_Error e =>
+ let '_ := (ext_handle_data_check_error e) : unit in
+ returnm RETIRE_FAIL
+ | Ext_DataAddr_OK vaddr =>
+ (if check_misaligned vaddr width then
+ (handle_mem_exception vaddr (E_SAMO_Addr_Align tt)) >> returnm RETIRE_FAIL
+ else
+ (translateAddr vaddr (Write Data)) >>= fun w__1 : TR_Result (mword 64) ExceptionType =>
+ (match w__1 with
+ | TR_Failure (e, _) => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ | TR_Address (addr, _) =>
+ (match width with
+ | BYTE => returnm (MemValue tt)
+ | HALF => returnm (MemValue tt)
+ | WORD => (mem_write_ea addr 4 aq rl false) : M (MemoryOpResult unit)
+ | DOUBLE => (mem_write_ea addr 8 aq rl false) : M (MemoryOpResult unit)
+ end) >>= fun eares : MemoryOpResult unit =>
+ (match eares with
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
+ | MemValue _ =>
+ (rF_bits rs2) >>= fun rs2_val =>
+ (match (width, 64) with
+ | (BYTE, g__320) => (handle_illegal tt) >> returnm RETIRE_FAIL
+ | (HALF, g__321) => (handle_illegal tt) >> returnm RETIRE_FAIL
+ | (WORD, g__322) =>
+ (mem_write_value addr 4 (subrange_vec_dec rs2_val 31 0) aq rl con) >>= fun w__4 : MemoryOpResult bool =>
+ (process_fstore vaddr w__4)
+ : M (Retired)
+ | (DOUBLE, l__9) =>
+ (if sumbool_of_bool (Z.eqb l__9 64) then
+ (mem_write_value addr 8 rs2_val aq rl con) >>= fun w__6 : MemoryOpResult bool =>
+ (process_fstore vaddr w__6)
+ : M (Retired)
+ else
+ assert_exp' false "Pattern match failure at model/riscv_insts_fext.sail 401:14 - 406:15" >>= fun _ =>
+ exit tt)
+ : M (Retired)
+ end)
+ : M (Retired)
+ end)
+ : M (Retired)
+ end)
+ : M (Retired))
+ : M (Retired)
+ end)
+ : M (Retired).
Definition execute_STORECON
-(aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (rd : mword 5)
+(aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (rd : mword 5)
: M (Retired) :=
-
(speculate_conditional tt) >>= fun w__0 : bool =>
- (if ((Bool.eqb w__0 false)) then
- (wX (projT1 (regidx_to_regno rd)) (EXTZ 64 (vec_of_bits [B1] : mword 1))) >>
- returnm (RETIRE_SUCCESS
- : Retired)
+ (if Bool.eqb w__0 false then (wX_bits rd (EXTZ 64 ('b"1" : mword 1))) >> returnm RETIRE_SUCCESS
else
(haveAtomics tt) >>= fun w__1 : bool =>
- (if sumbool_of_bool (w__1) then
- (ext_data_get_addr rs1 (zeros_implicit 64) Write width) >>= fun w__2 : Ext_DataAddr_Check unit =>
+ (if sumbool_of_bool w__1 then
+ (ext_data_get_addr rs1 (zeros_implicit 64) (Write Data) width) >>= fun w__2 : Ext_DataAddr_Check unit =>
(match w__2 with
- | Ext_DataAddr_Error (e) =>
+ | Ext_DataAddr_Error e =>
let '_ := (ext_handle_data_check_error e) : unit in
- returnm (RETIRE_FAIL
- : Retired)
- | Ext_DataAddr_OK (vaddr) =>
- (match width with
- | BYTE =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | HALF =>
- (cast_unit_vec (access_vec_dec vaddr 0)) >>= fun w__3 : mword 1 =>
- returnm ((eq_vec w__3 (vec_of_bits [B0] : mword 1))
- : bool)
- | WORD =>
- returnm ((eq_vec (subrange_vec_dec vaddr 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))
- : bool)
- | DOUBLE =>
- returnm ((eq_vec (subrange_vec_dec vaddr 2 0)
- (vec_of_bits [B0;B0;B0] : mword (2 - 0 + 1)))
- : bool)
- end) >>= fun aligned : bool =>
- (if sumbool_of_bool ((negb aligned)) then
- (handle_mem_exception vaddr E_SAMO_Addr_Align) >> returnm (RETIRE_FAIL : Retired)
- else if ((Bool.eqb (match_reservation vaddr) false)) then
- (wX (projT1 (regidx_to_regno rd)) (EXTZ 64 (vec_of_bits [B1] : mword 1))) >>
+ returnm RETIRE_FAIL
+ | Ext_DataAddr_OK vaddr =>
+ let aligned : bool :=
+ match width with
+ | BYTE => true
+ | HALF => eq_vec (subrange_vec_dec vaddr 0 0) ('b"0" : mword (0 - 0 + 1))
+ | WORD => eq_vec (subrange_vec_dec vaddr 1 0) ('b"00" : mword (1 - 0 + 1))
+ | DOUBLE => eq_vec (subrange_vec_dec vaddr 2 0) ('b"000" : mword (2 - 0 + 1))
+ end in
+ (if sumbool_of_bool (negb aligned) then
+ (handle_mem_exception vaddr (E_SAMO_Addr_Align tt)) >> returnm RETIRE_FAIL
+ else if Bool.eqb (match_reservation vaddr) false then
+ (wX_bits rd (EXTZ 64 ('b"1" : mword 1))) >>
let '_ := (cancel_reservation tt) : unit in
- returnm (RETIRE_SUCCESS
- : Retired)
+ returnm RETIRE_SUCCESS
else
- (translateAddr vaddr Write) >>= fun w__4 : TR_Result (mword 64) ExceptionType =>
- (match w__4 with
- | TR_Failure (e) =>
- (handle_mem_exception vaddr e) >> returnm (RETIRE_FAIL : Retired)
- | TR_Address (addr) =>
+ (translateAddr vaddr (Write Data)) >>= fun w__3 : TR_Result (mword 64) ExceptionType =>
+ (match w__3 with
+ | TR_Failure (e, _) => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ | TR_Address (addr, _) =>
(match (width, 64) with
- | (WORD, g__217) => (mem_write_ea addr 4 aq rl true) : M (MemoryOpResult unit)
+ | (WORD, g__313) => (mem_write_ea addr 4 aq rl true) : M (MemoryOpResult unit)
| (DOUBLE, l__4) =>
- (if sumbool_of_bool ((Z.eqb l__4 64)) then
+ (if sumbool_of_bool (Z.eqb l__4 64) then
(mem_write_ea addr 8 aq rl true)
: M (MemoryOpResult unit)
else
@@ -22062,16 +28814,15 @@ Definition execute_STORECON
: M (MemoryOpResult unit)
end) >>= fun eares : MemoryOpResult unit =>
(match eares with
- | MemException (e) =>
- (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
- | MemValue (_) =>
- (rX (projT1 (regidx_to_regno rs2))) >>= fun rs2_val =>
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
+ | MemValue _ =>
+ (rX_bits rs2) >>= fun rs2_val =>
(match (width, 64) with
- | (WORD, g__216) =>
+ | (WORD, g__312) =>
(mem_write_value addr 4 (subrange_vec_dec rs2_val 31 0) aq rl true)
: M (MemoryOpResult bool)
| (DOUBLE, l__3) =>
- (if sumbool_of_bool ((Z.eqb l__3 64)) then
+ (if sumbool_of_bool (Z.eqb l__3 64) then
(mem_write_value addr 8 rs2_val aq rl true)
: M (MemoryOpResult bool)
else
@@ -22083,20 +28834,15 @@ Definition execute_STORECON
: M (MemoryOpResult bool)
end) >>= fun res : MemoryOpResult bool =>
(match res with
- | MemValue (true) =>
- (wX (projT1 (regidx_to_regno rd))
- (EXTZ 64 (vec_of_bits [B0] : mword 1))) >>
+ | MemValue true =>
+ (wX_bits rd (EXTZ 64 ('b"0" : mword 1))) >>
let '_ := (cancel_reservation tt) : unit in
- returnm (RETIRE_SUCCESS
- : Retired)
- | MemValue (false) =>
- (wX (projT1 (regidx_to_regno rd))
- (EXTZ 64 (vec_of_bits [B1] : mword 1))) >>
+ returnm RETIRE_SUCCESS
+ | MemValue false =>
+ (wX_bits rd (EXTZ 64 ('b"1" : mword 1))) >>
let '_ := (cancel_reservation tt) : unit in
- returnm (RETIRE_SUCCESS
- : Retired)
- | MemException (e) =>
- (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
+ returnm RETIRE_SUCCESS
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
end)
: M (Retired)
end)
@@ -22106,54 +28852,50 @@ Definition execute_STORECON
: M (Retired)
end)
: M (Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired))
: M (Retired).
Definition execute_STORE
-(imm : mword 12) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (aq : bool) (rl : bool)
+(imm : mword 12) (rs2 : mword 5) (rs1 : mword 5) (width : word_width) (aq : bool) (rl : bool)
: M (Retired) :=
-
let offset : xlenbits := EXTS 64 imm in
- (ext_data_get_addr rs1 offset Write width) >>= fun w__0 : Ext_DataAddr_Check unit =>
+ (ext_data_get_addr rs1 offset (Write Data) width) >>= fun w__0 : Ext_DataAddr_Check unit =>
(match w__0 with
- | Ext_DataAddr_Error (e) =>
+ | Ext_DataAddr_Error e =>
let '_ := (ext_handle_data_check_error e) : unit in
- returnm (RETIRE_FAIL
- : Retired)
- | Ext_DataAddr_OK (vaddr) =>
- (check_misaligned vaddr width) >>= fun w__1 : bool =>
- (if sumbool_of_bool (w__1) then
- (handle_mem_exception vaddr E_SAMO_Addr_Align) >> returnm (RETIRE_FAIL : Retired)
+ returnm RETIRE_FAIL
+ | Ext_DataAddr_OK vaddr =>
+ (if check_misaligned vaddr width then
+ (handle_mem_exception vaddr (E_SAMO_Addr_Align tt)) >> returnm RETIRE_FAIL
else
- (translateAddr vaddr Write) >>= fun w__2 : TR_Result (mword 64) ExceptionType =>
- (match w__2 with
- | TR_Failure (e) => (handle_mem_exception vaddr e) >> returnm (RETIRE_FAIL : Retired)
- | TR_Address (addr) =>
+ (translateAddr vaddr (Write Data)) >>= fun w__1 : TR_Result (mword 64) ExceptionType =>
+ (match w__1 with
+ | TR_Failure (e, _) => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ | TR_Address (paddr, _) =>
(match width with
- | BYTE => (mem_write_ea addr 1 aq rl false) : M (MemoryOpResult unit)
- | HALF => (mem_write_ea addr 2 aq rl false) : M (MemoryOpResult unit)
- | WORD => (mem_write_ea addr 4 aq rl false) : M (MemoryOpResult unit)
- | DOUBLE => (mem_write_ea addr 8 aq rl false) : M (MemoryOpResult unit)
+ | BYTE => (mem_write_ea paddr 1 aq rl false) : M (MemoryOpResult unit)
+ | HALF => (mem_write_ea paddr 2 aq rl false) : M (MemoryOpResult unit)
+ | WORD => (mem_write_ea paddr 4 aq rl false) : M (MemoryOpResult unit)
+ | DOUBLE => (mem_write_ea paddr 8 aq rl false) : M (MemoryOpResult unit)
end) >>= fun eares : MemoryOpResult unit =>
(match eares with
- | MemException (e) =>
- (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
- | MemValue (_) =>
- (rX (projT1 (regidx_to_regno rs2))) >>= fun rs2_val =>
+ | MemException e => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ | MemValue _ =>
+ (rX_bits rs2) >>= fun rs2_val =>
(match (width, 64) with
- | (BYTE, g__210) =>
- (mem_write_value addr 1 (subrange_vec_dec rs2_val 7 0) aq rl false)
+ | (BYTE, g__306) =>
+ (mem_write_value paddr 1 (subrange_vec_dec rs2_val 7 0) aq rl false)
: M (MemoryOpResult bool)
- | (HALF, g__211) =>
- (mem_write_value addr 2 (subrange_vec_dec rs2_val 15 0) aq rl false)
+ | (HALF, g__307) =>
+ (mem_write_value paddr 2 (subrange_vec_dec rs2_val 15 0) aq rl false)
: M (MemoryOpResult bool)
- | (WORD, g__212) =>
- (mem_write_value addr 4 (subrange_vec_dec rs2_val 31 0) aq rl false)
+ | (WORD, g__308) =>
+ (mem_write_value paddr 4 (subrange_vec_dec rs2_val 31 0) aq rl false)
: M (MemoryOpResult bool)
| (DOUBLE, l__1) =>
- (if sumbool_of_bool ((Z.eqb l__1 64)) then
- (mem_write_value addr 8 rs2_val aq rl false)
+ (if sumbool_of_bool (Z.eqb l__1 64) then
+ (mem_write_value paddr 8 rs2_val aq rl false)
: M (MemoryOpResult bool)
else
assert_exp' false "Pattern match failure at model/riscv_insts_base.sail 394:47 - 399:15" >>= fun _ =>
@@ -22161,11 +28903,10 @@ Definition execute_STORE
: M (MemoryOpResult bool)
end) >>= fun res : MemoryOpResult bool =>
(match res with
- | MemValue (true) => returnm (RETIRE_SUCCESS : Retired)
- | MemValue (false) =>
+ | MemValue true => returnm RETIRE_SUCCESS
+ | MemValue false =>
(internal_error "store got false from mem_write_value") : M (Retired)
- | MemException (e) =>
- (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
+ | MemException e => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
end)
: M (Retired)
end)
@@ -22176,43 +28917,29 @@ Definition execute_STORE
end)
: M (Retired).
-Definition execute_SRET '(tt : unit)
-: M (Retired) :=
-
+Definition execute_SRET '(tt : unit) : M (Retired) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
(match w__0 with
- | User => (handle_illegal tt) : M (unit)
+ | User => returnm true
| Supervisor =>
(or_boolM ((haveSupMode tt) >>= fun w__1 : bool => returnm ((negb w__1) : bool))
(read_reg mstatus_ref >>= fun w__2 : Mstatus =>
- returnm ((eq_vec (_get_Mstatus_TSR w__2) ((bool_to_bits true) : mword 1))
- : bool))) >>= fun w__3 : bool =>
- (if sumbool_of_bool (w__3) then (handle_illegal tt) : M (unit)
- else
- read_reg cur_privilege_ref >>= fun w__4 : Privilege =>
- ((read_reg PC_ref) : M (mword 64)) >>= fun w__5 : mword 64 =>
- (exception_handler w__4 (CTL_SRET (tt)) w__5) >>= fun w__6 : mword 64 =>
- (set_next_pc w__6)
- : M (unit))
- : M (unit)
- | Machine =>
- (haveSupMode tt) >>= fun w__7 : bool =>
- (if sumbool_of_bool ((negb w__7)) then (handle_illegal tt) : M (unit)
- else
- read_reg cur_privilege_ref >>= fun w__8 : Privilege =>
- ((read_reg PC_ref) : M (mword 64)) >>= fun w__9 : mword 64 =>
- (exception_handler w__8 (CTL_SRET (tt)) w__9) >>= fun w__10 : mword 64 =>
- (set_next_pc w__10)
- : M (unit))
- : M (unit)
- end) >>
- returnm (RETIRE_FAIL
- : Retired).
+ returnm ((eq_vec (_get_Mstatus_TSR w__2) ('b"1" : mword 1)) : bool)))
+ : M (bool)
+ | Machine => (haveSupMode tt) >>= fun w__4 : bool => returnm (negb w__4)
+ end) >>= fun sret_illegal : bool =>
+ (if sumbool_of_bool sret_illegal then (handle_illegal tt) : M (unit)
+ else if negb (ext_check_xret_priv Supervisor) then returnm (ext_fail_xret_priv tt)
+ else
+ read_reg cur_privilege_ref >>= fun w__5 : Privilege =>
+ ((read_reg PC_ref) : M (mword 64)) >>= fun w__6 : mword 64 =>
+ (exception_handler w__5 (CTL_SRET tt) w__6) >>= fun w__7 : mword 64 =>
+ (set_next_pc w__7)
+ : M (unit)) >>
+ returnm RETIRE_FAIL.
-Definition execute_SHIFTW (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : sop)
-: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun w__0 : mword 64 =>
+Definition execute_SHIFTW (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : sop) : M (Retired) :=
+ (rX_bits rs1) >>= fun w__0 : mword 64 =>
let rs1_val := subrange_vec_dec w__0 31 0 in
let result : bits 32 :=
match op with
@@ -22220,67 +28947,49 @@ Definition execute_SHIFTW (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op :
| RISCV_SRLI => shift_bits_right rs1_val shamt
| RISCV_SRAI => shift_right_arith32 rs1_val shamt
end in
- (wX (projT1 (regidx_to_regno rd)) (EXTS 64 result)) >> returnm (RETIRE_SUCCESS : Retired).
+ (wX_bits rd (EXTS 64 result)) >> returnm RETIRE_SUCCESS.
-Definition execute_SHIFTIWOP (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : sopw)
+Definition execute_SHIFTIWOP (shamt : mword 5) (rs1 : mword 5) (rd : mword 5) (op : sopw)
: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun rs1_val =>
+ (rX_bits rs1) >>= fun rs1_val =>
let result : bits 32 :=
match op with
| RISCV_SLLIW => shift_bits_left (subrange_vec_dec rs1_val 31 0) shamt
| RISCV_SRLIW => shift_bits_right (subrange_vec_dec rs1_val 31 0) shamt
| RISCV_SRAIW => shift_right_arith32 (subrange_vec_dec rs1_val 31 0) shamt
end in
- (wX (projT1 (regidx_to_regno rd)) (EXTS 64 result)) >> returnm (RETIRE_SUCCESS : Retired).
+ (wX_bits rd (EXTS 64 result)) >> returnm RETIRE_SUCCESS.
-Definition execute_SHIFTIOP (shamt : mword 6) (rs1 : mword 5) (rd : mword 5) (op : sop)
+Definition execute_SHIFTIOP (shamt : mword 6) (rs1 : mword 5) (rd : mword 5) (op : sop)
: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun rs1_val =>
+ (rX_bits rs1) >>= fun rs1_val =>
let result : xlenbits :=
match op with
| RISCV_SLLI => shift_bits_left rs1_val shamt
| RISCV_SRLI => shift_bits_right rs1_val shamt
| RISCV_SRAI => shift_right_arith64 rs1_val shamt
end in
- (wX (projT1 (regidx_to_regno rd)) result) >> returnm (RETIRE_SUCCESS : Retired).
+ (wX_bits rd result) >> returnm RETIRE_SUCCESS.
-Definition execute_SFENCE_VMA (rs1 : mword 5) (rs2 : mword 5)
-: M (Retired) :=
-
- (if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then
- returnm (None
- : option (mword 64))
- else
- (rX (projT1 (regidx_to_regno rs1))) >>= fun w__0 : mword 64 =>
- returnm ((Some
- (w__0))
- : option (mword 64))) >>= fun addr : option xlenbits =>
- (if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs2)) 0)) then
- returnm (None
- : option (mword 64))
- else
- (rX (projT1 (regidx_to_regno rs2))) >>= fun w__1 : mword 64 =>
- returnm ((Some
- (w__1))
- : option (mword 64))) >>= fun asid : option xlenbits =>
+Definition execute_SFENCE_VMA (rs1 : mword 5) (rs2 : mword 5) : M (Retired) :=
+ (if eq_vec rs1 ('b"00000" : mword 5) then returnm None
+ else (rX_bits rs1) >>= fun w__0 : mword 64 => returnm (Some w__0)) >>= fun addr : option xlenbits =>
+ (if eq_vec rs2 ('b"00000" : mword 5) then returnm None
+ else (rX_bits rs2) >>= fun w__1 : mword 64 => returnm (Some w__1)) >>= fun asid : option xlenbits =>
read_reg cur_privilege_ref >>= fun w__2 : Privilege =>
(match w__2 with
- | User => (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired)
+ | User => (handle_illegal tt) >> returnm RETIRE_FAIL
| Supervisor =>
read_reg mstatus_ref >>= fun w__3 : Mstatus =>
read_reg mstatus_ref >>= fun w__4 : Mstatus =>
- let p__214 := (architecture (get_mstatus_SXL w__3), _get_Mstatus_TVM w__4) in
- (match p__214 with
- | (Some (g__213), v_0) =>
- (if ((eq_vec v_0 ((bool_to_bits true) : mword 1))) then
- (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired)
- else if ((eq_vec v_0 ((bool_to_bits false) : mword 1))) then
- (flush_TLB asid addr) >> returnm (RETIRE_SUCCESS : Retired)
+ let p__310 := (architecture (get_mstatus_SXL w__3), _get_Mstatus_TVM w__4) in
+ (match p__310 with
+ | (Some g__309, b__0) =>
+ (if eq_vec b__0 ('b"1" : mword 1) then (handle_illegal tt) >> returnm RETIRE_FAIL
+ else if eq_vec b__0 ('b"0" : mword 1) then
+ (flush_TLB asid addr) >> returnm RETIRE_SUCCESS
else
- (match (Some
- (g__213), v_0) with
+ (match (Some g__309, b__0) with
| (_, _) => (internal_error "unimplemented sfence architecture") : M (Retired)
end)
: M (Retired))
@@ -22288,16 +28997,14 @@ Definition execute_SFENCE_VMA (rs1 : mword 5) (rs2 : mword 5)
| (_, _) => (internal_error "unimplemented sfence architecture") : M (Retired)
end)
: M (Retired)
- | Machine => (flush_TLB asid addr) >> returnm (RETIRE_SUCCESS : Retired)
+ | Machine => (flush_TLB asid addr) >> returnm RETIRE_SUCCESS
end)
: M (Retired).
-Definition execute_RTYPEW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : ropw)
-: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun w__0 : mword 64 =>
+Definition execute_RTYPEW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : ropw) : M (Retired) :=
+ (rX_bits rs1) >>= fun w__0 : mword 64 =>
let rs1_val := subrange_vec_dec w__0 31 0 in
- (rX (projT1 (regidx_to_regno rs2))) >>= fun w__1 : mword 64 =>
+ (rX_bits rs2) >>= fun w__1 : mword 64 =>
let rs2_val := subrange_vec_dec w__1 31 0 in
let result : bits 32 :=
match op with
@@ -22307,13 +29014,11 @@ Definition execute_RTYPEW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : r
| RISCV_SRLW => shift_bits_right rs1_val (subrange_vec_dec rs2_val 4 0)
| RISCV_SRAW => shift_right_arith32 rs1_val (subrange_vec_dec rs2_val 4 0)
end in
- (wX (projT1 (regidx_to_regno rd)) (EXTS 64 result)) >> returnm (RETIRE_SUCCESS : Retired).
+ (wX_bits rd (EXTS 64 result)) >> returnm RETIRE_SUCCESS.
-Definition execute_RTYPE (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : rop)
-: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun rs1_val =>
- (rX (projT1 (regidx_to_regno rs2))) >>= fun rs2_val =>
+Definition execute_RTYPE (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : rop) : M (Retired) :=
+ (rX_bits rs1) >>= fun rs1_val =>
+ (rX_bits rs2) >>= fun rs2_val =>
let result : xlenbits :=
match op with
| RISCV_ADD => add_vec rs1_val rs2_val
@@ -22327,192 +29032,196 @@ Definition execute_RTYPE (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (op : ro
| RISCV_SUB => sub_vec rs1_val rs2_val
| RISCV_SRA => shift_right_arith64 rs1_val (subrange_vec_dec rs2_val 5 0)
end in
- (wX (projT1 (regidx_to_regno rd)) result) >> returnm (RETIRE_SUCCESS : Retired).
+ (wX_bits rd result) >> returnm RETIRE_SUCCESS.
-Definition execute_RISCV_JALR (imm : mword 12) (rs1 : mword 5) (rd : mword 5)
-: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun w__0 : mword 64 =>
+Definition execute_RISCV_JALR (imm : mword 12) (rs1 : mword 5) (rd : mword 5) : M (Retired) :=
+ (rX_bits rs1) >>= fun w__0 : mword 64 =>
let t : xlenbits := add_vec w__0 (EXTS 64 imm) in
(match (ext_control_check_addr t) with
- | Ext_ControlAddr_Error (e) =>
+ | Ext_ControlAddr_Error e =>
let '_ := (ext_handle_control_check_error e) : unit in
- returnm (RETIRE_FAIL
- : Retired)
- | Ext_ControlAddr_OK (addr) =>
+ returnm RETIRE_FAIL
+ | Ext_ControlAddr_OK addr =>
let target := update_vec_dec addr 0 B0 in
(and_boolM ((bit_to_bool (access_vec_dec target 1)) : M (bool))
((haveRVC tt) >>= fun w__2 : bool => returnm ((negb w__2) : bool))) >>= fun w__3 : bool =>
- (if sumbool_of_bool (w__3) then
- (handle_mem_exception target E_Fetch_Addr_Align) >> returnm (RETIRE_FAIL : Retired)
+ (if sumbool_of_bool w__3 then
+ (handle_mem_exception target (E_Fetch_Addr_Align tt)) >> returnm RETIRE_FAIL
else
(get_next_pc tt) >>= fun w__4 : mword 64 =>
- (wX (projT1 (regidx_to_regno rd)) w__4) >>
- (set_next_pc target) >> returnm (RETIRE_SUCCESS : Retired))
+ (wX_bits rd w__4) >> (set_next_pc target) >> returnm RETIRE_SUCCESS)
: M (Retired)
end)
: M (Retired).
-Definition execute_RISCV_JAL (imm : mword 21) (rd : mword 5)
-: M (Retired) :=
-
+Definition execute_RISCV_JAL (imm : mword 21) (rd : mword 5) : M (Retired) :=
((read_reg PC_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
let t : xlenbits := add_vec w__0 (EXTS 64 imm) in
(match (ext_control_check_pc t) with
- | Ext_ControlAddr_Error (e) =>
+ | Ext_ControlAddr_Error e =>
let '_ := (ext_handle_control_check_error e) : unit in
- returnm (RETIRE_FAIL
- : Retired)
- | Ext_ControlAddr_OK (target) =>
+ returnm RETIRE_FAIL
+ | Ext_ControlAddr_OK target =>
(and_boolM ((bit_to_bool (access_vec_dec target 1)) : M (bool))
((haveRVC tt) >>= fun w__2 : bool => returnm ((negb w__2) : bool))) >>= fun w__3 : bool =>
- (if sumbool_of_bool (w__3) then
- (handle_mem_exception target E_Fetch_Addr_Align) >> returnm (RETIRE_FAIL : Retired)
+ (if sumbool_of_bool w__3 then
+ (handle_mem_exception target (E_Fetch_Addr_Align tt)) >> returnm RETIRE_FAIL
else
(get_next_pc tt) >>= fun w__4 : mword 64 =>
- (wX (projT1 (regidx_to_regno rd)) w__4) >>
- (set_next_pc target) >> returnm (RETIRE_SUCCESS : Retired))
+ (wX_bits rd w__4) >> (set_next_pc target) >> returnm RETIRE_SUCCESS)
: M (Retired)
end)
: M (Retired).
-Definition execute_REMW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
-: M (Retired) :=
-
+Definition execute_REMW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool) : M (Retired) :=
(haveMulDiv tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool (w__0) then
- (rX (projT1 (regidx_to_regno rs1))) >>= fun w__1 : mword 64 =>
+ (if sumbool_of_bool w__0 then
+ (rX_bits rs1) >>= fun w__1 : mword 64 =>
let rs1_val := subrange_vec_dec w__1 31 0 in
- (rX (projT1 (regidx_to_regno rs2))) >>= fun w__2 : mword 64 =>
+ (rX_bits rs2) >>= fun w__2 : mword 64 =>
let rs2_val := subrange_vec_dec w__2 31 0 in
- let rs1_int : Z :=
- if sumbool_of_bool (s) then projT1 (sint rs1_val)
- else projT1 (uint rs1_val) in
- let rs2_int : Z :=
- if sumbool_of_bool (s) then projT1 (sint rs2_val)
- else projT1 (uint rs2_val) in
- let r : Z := if sumbool_of_bool ((Z.eqb rs2_int 0)) then rs1_int else Z.rem rs1_int rs2_int in
- (wX (projT1 (regidx_to_regno rd)) (EXTS 64 (to_bits 32 r))) >>
- returnm (RETIRE_SUCCESS
- : Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ let rs1_int : Z := if sumbool_of_bool s then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
+ let rs2_int : Z := if sumbool_of_bool s then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
+ let r : Z := if sumbool_of_bool (Z.eqb rs2_int 0) then rs1_int else Z.rem rs1_int rs2_int in
+ (wX_bits rd (EXTS 64 (to_bits 32 r))) >> returnm RETIRE_SUCCESS
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired).
-Definition execute_REM (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
-: M (Retired) :=
-
+Definition execute_REM (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool) : M (Retired) :=
(haveMulDiv tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool (w__0) then
- (rX (projT1 (regidx_to_regno rs1))) >>= fun rs1_val =>
- (rX (projT1 (regidx_to_regno rs2))) >>= fun rs2_val =>
- let rs1_int : Z :=
- if sumbool_of_bool (s) then projT1 (sint rs1_val)
- else projT1 (uint rs1_val) in
- let rs2_int : Z :=
- if sumbool_of_bool (s) then projT1 (sint rs2_val)
- else projT1 (uint rs2_val) in
- let r : Z := if sumbool_of_bool ((Z.eqb rs2_int 0)) then rs1_int else Z.rem rs1_int rs2_int in
- (wX (projT1 (regidx_to_regno rd)) (to_bits 64 r)) >> returnm (RETIRE_SUCCESS : Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ (if sumbool_of_bool w__0 then
+ (rX_bits rs1) >>= fun rs1_val =>
+ (rX_bits rs2) >>= fun rs2_val =>
+ let rs1_int : Z := if sumbool_of_bool s then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
+ let rs2_int : Z := if sumbool_of_bool s then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
+ let r : Z := if sumbool_of_bool (Z.eqb rs2_int 0) then rs1_int else Z.rem rs1_int rs2_int in
+ (wX_bits rd (to_bits 64 r)) >> returnm RETIRE_SUCCESS
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired).
-Definition execute_MULW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5)
-: M (Retired) :=
-
+Definition execute_MULW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) : M (Retired) :=
(haveMulDiv tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool (w__0) then
- (rX (projT1 (regidx_to_regno rs1))) >>= fun w__1 : mword 64 =>
+ (if sumbool_of_bool w__0 then
+ (rX_bits rs1) >>= fun w__1 : mword 64 =>
let rs1_val := subrange_vec_dec w__1 31 0 in
- (rX (projT1 (regidx_to_regno rs2))) >>= fun w__2 : mword 64 =>
+ (rX_bits rs2) >>= fun w__2 : mword 64 =>
let rs2_val := subrange_vec_dec w__2 31 0 in
let rs1_int : Z := projT1 (sint rs1_val) in
let rs2_int : Z := projT1 (sint rs2_val) in
let result32 := subrange_vec_dec (to_bits 64 (Z.mul rs1_int rs2_int)) 31 0 in
let result : xlenbits := EXTS 64 result32 in
- (wX (projT1 (regidx_to_regno rd)) result) >> returnm (RETIRE_SUCCESS : Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ (wX_bits rd result) >> returnm RETIRE_SUCCESS
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired).
Definition execute_MUL
-(rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (high : bool) (signed1 : bool) (signed2 : bool)
+(rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (high : bool) (signed1 : bool) (signed2 : bool)
: M (Retired) :=
-
(haveMulDiv tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool (w__0) then
- (rX (projT1 (regidx_to_regno rs1))) >>= fun rs1_val =>
- (rX (projT1 (regidx_to_regno rs2))) >>= fun rs2_val =>
+ (if sumbool_of_bool w__0 then
+ (rX_bits rs1) >>= fun rs1_val =>
+ (rX_bits rs2) >>= fun rs2_val =>
let rs1_int : Z :=
- if sumbool_of_bool (signed1) then projT1 (sint rs1_val)
+ if sumbool_of_bool signed1 then projT1 (sint rs1_val)
else projT1 (uint rs1_val) in
let rs2_int : Z :=
- if sumbool_of_bool (signed2) then projT1 (sint rs2_val)
+ if sumbool_of_bool signed2 then projT1 (sint rs2_val)
else projT1 (uint rs2_val) in
let result_wide := to_bits (Z.mul 2 64) (Z.mul rs1_int rs2_int) in
let result :=
- if sumbool_of_bool (high) then subrange_vec_dec result_wide (Z.sub (Z.mul 2 64) 1) 64
+ if sumbool_of_bool high then subrange_vec_dec result_wide (Z.sub (Z.mul 2 64) 1) 64
else subrange_vec_dec result_wide (Z.sub 64 1) 0 in
- (wX (projT1 (regidx_to_regno rd)) result) >> returnm (RETIRE_SUCCESS : Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ (wX_bits rd result) >> returnm RETIRE_SUCCESS
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired).
-Definition execute_MRET '(tt : unit)
-: M (Retired) :=
-
+Definition execute_MRET '(tt : unit) : M (Retired) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
- (if ((eq_vec (privLevel_to_bits w__0) ((privLevel_to_bits Machine) : mword 2))) then
+ (if generic_neq w__0 Machine then (handle_illegal tt) : M (unit)
+ else if negb (ext_check_xret_priv Machine) then returnm (ext_fail_xret_priv tt)
+ else
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
((read_reg PC_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
- (exception_handler w__1 (CTL_MRET (tt)) w__2) >>= fun w__3 : mword 64 =>
+ (exception_handler w__1 (CTL_MRET tt) w__2) >>= fun w__3 : mword 64 =>
(set_next_pc w__3)
- : M (unit)
- else (handle_illegal tt) : M (unit)) >>
- returnm (RETIRE_FAIL
- : Retired).
+ : M (unit)) >>
+ returnm RETIRE_FAIL.
+
+Definition execute_LOAD_FP (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (width : word_width)
+: M (Retired) :=
+ let offset : xlenbits := EXTS 64 imm in
+ (ext_data_get_addr rs1 offset (Read Data) width) >>= fun w__0 : Ext_DataAddr_Check unit =>
+ (match w__0 with
+ | Ext_DataAddr_Error e =>
+ let '_ := (ext_handle_data_check_error e) : unit in
+ returnm RETIRE_FAIL
+ | Ext_DataAddr_OK vaddr =>
+ (if check_misaligned vaddr width then
+ (handle_mem_exception vaddr (E_Load_Addr_Align tt)) >> returnm RETIRE_FAIL
+ else
+ (translateAddr vaddr (Read Data)) >>= fun w__1 : TR_Result (mword 64) ExceptionType =>
+ (match w__1 with
+ | TR_Failure (e, _) => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ | TR_Address (addr, _) =>
+ let '(aq, rl, res) := (false, false, false) in
+ (match (width, 64) with
+ | (BYTE, g__317) => (handle_illegal tt) >> returnm RETIRE_FAIL
+ | (HALF, g__318) => (handle_illegal tt) >> returnm RETIRE_FAIL
+ | (WORD, g__319) =>
+ (mem_read (Read Data) addr 4 aq rl res) >>= fun w__2 : MemoryOpResult (mword (8 * 4)) =>
+ (process_fload32 rd vaddr w__2)
+ : M (Retired)
+ | (DOUBLE, l__8) =>
+ (if sumbool_of_bool (Z.eqb l__8 64) then
+ (mem_read (Read Data) addr 8 aq rl res) >>= fun w__4 : MemoryOpResult (mword (8 * 8)) =>
+ (process_fload64 rd vaddr w__4)
+ : M (Retired)
+ else
+ assert_exp' false "Pattern match failure at model/riscv_insts_fext.sail 331:10 - 338:11" >>= fun _ =>
+ exit tt)
+ : M (Retired)
+ end)
+ : M (Retired)
+ end)
+ : M (Retired))
+ : M (Retired)
+ end)
+ : M (Retired).
Definition execute_LOADRES
-(aq : bool) (rl : bool) (rs1 : mword 5) (width : word_width) (rd : mword 5)
+(aq : bool) (rl : bool) (rs1 : mword 5) (width : word_width) (rd : mword 5)
: M (Retired) :=
-
(haveAtomics tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool (w__0) then
- (ext_data_get_addr rs1 (zeros_implicit 64) Read width) >>= fun w__1 : Ext_DataAddr_Check unit =>
+ (if sumbool_of_bool w__0 then
+ (ext_data_get_addr rs1 (zeros_implicit 64) (Read Data) width) >>= fun w__1 : Ext_DataAddr_Check unit =>
(match w__1 with
- | Ext_DataAddr_Error (e) =>
+ | Ext_DataAddr_Error e =>
let '_ := (ext_handle_data_check_error e) : unit in
- returnm (RETIRE_FAIL
- : Retired)
- | Ext_DataAddr_OK (vaddr) =>
- (match width with
- | BYTE => returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | HALF =>
- (cast_unit_vec (access_vec_dec vaddr 0)) >>= fun w__2 : mword 1 =>
- returnm ((eq_vec w__2 (vec_of_bits [B0] : mword 1))
- : bool)
- | WORD =>
- returnm ((eq_vec (subrange_vec_dec vaddr 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))
- : bool)
- | DOUBLE =>
- returnm ((eq_vec (subrange_vec_dec vaddr 2 0)
- (vec_of_bits [B0;B0;B0] : mword (2 - 0 + 1)))
- : bool)
- end) >>= fun aligned : bool =>
- (if sumbool_of_bool ((negb aligned)) then
- (handle_mem_exception vaddr E_Load_Addr_Align) >> returnm (RETIRE_FAIL : Retired)
+ returnm RETIRE_FAIL
+ | Ext_DataAddr_OK vaddr =>
+ let aligned : bool :=
+ match width with
+ | BYTE => true
+ | HALF => eq_vec (subrange_vec_dec vaddr 0 0) ('b"0" : mword (0 - 0 + 1))
+ | WORD => eq_vec (subrange_vec_dec vaddr 1 0) ('b"00" : mword (1 - 0 + 1))
+ | DOUBLE => eq_vec (subrange_vec_dec vaddr 2 0) ('b"000" : mword (2 - 0 + 1))
+ end in
+ (if sumbool_of_bool (negb aligned) then
+ (handle_mem_exception vaddr (E_Load_Addr_Align tt)) >> returnm RETIRE_FAIL
else
- (translateAddr vaddr Read) >>= fun w__3 : TR_Result (mword 64) ExceptionType =>
- (match w__3 with
- | TR_Failure (e) => (handle_mem_exception vaddr e) >> returnm (RETIRE_FAIL : Retired)
- | TR_Address (addr) =>
+ (translateAddr vaddr (Read Data)) >>= fun w__2 : TR_Result (mword 64) ExceptionType =>
+ (match w__2 with
+ | TR_Failure (e, _) => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ | TR_Address (addr, _) =>
(match (width, 64) with
- | (WORD, g__215) =>
- (mem_read Read addr 4 aq rl true) >>= fun w__4 : MemoryOpResult (mword (8 * 4)) =>
- (process_loadres rd vaddr w__4 false)
+ | (WORD, g__311) =>
+ (mem_read (Read Data) addr 4 aq rl true) >>= fun w__3 : MemoryOpResult (mword (8 * 4)) =>
+ (process_loadres rd vaddr w__3 false)
: M (Retired)
| (DOUBLE, l__2) =>
- (if sumbool_of_bool ((Z.eqb l__2 64)) then
- (mem_read Read addr 8 aq rl true) >>= fun w__6 : MemoryOpResult (mword (8 * 8)) =>
- (process_loadres rd vaddr w__6 false)
+ (if sumbool_of_bool (Z.eqb l__2 64) then
+ (mem_read (Read Data) addr 8 aq rl true) >>= fun w__5 : MemoryOpResult (mword (8 * 8)) =>
+ (process_loadres rd vaddr w__5 false)
: M (Retired)
else (internal_error "LOADRES expected WORD or DOUBLE") : M (Retired))
: M (Retired)
@@ -22524,47 +29233,44 @@ Definition execute_LOADRES
: M (Retired)
end)
: M (Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired).
Definition execute_LOAD
(imm : mword 12) (rs1 : mword 5) (rd : mword 5) (is_unsigned : bool) (width : word_width)
-(aq : bool) (rl : bool)
+(aq : bool) (rl : bool)
: M (Retired) :=
-
let offset : xlenbits := EXTS 64 imm in
- (ext_data_get_addr rs1 offset Read width) >>= fun w__0 : Ext_DataAddr_Check unit =>
+ (ext_data_get_addr rs1 offset (Read Data) width) >>= fun w__0 : Ext_DataAddr_Check unit =>
(match w__0 with
- | Ext_DataAddr_Error (e) =>
+ | Ext_DataAddr_Error e =>
let '_ := (ext_handle_data_check_error e) : unit in
- returnm (RETIRE_FAIL
- : Retired)
- | Ext_DataAddr_OK (vaddr) =>
- (check_misaligned vaddr width) >>= fun w__1 : bool =>
- (if sumbool_of_bool (w__1) then
- (handle_mem_exception vaddr E_Load_Addr_Align) >> returnm (RETIRE_FAIL : Retired)
+ returnm RETIRE_FAIL
+ | Ext_DataAddr_OK vaddr =>
+ (if check_misaligned vaddr width then
+ (handle_mem_exception vaddr (E_Load_Addr_Align tt)) >> returnm RETIRE_FAIL
else
- (translateAddr vaddr Read) >>= fun w__2 : TR_Result (mword 64) ExceptionType =>
- (match w__2 with
- | TR_Failure (e) => (handle_mem_exception vaddr e) >> returnm (RETIRE_FAIL : Retired)
- | TR_Address (addr) =>
+ (translateAddr vaddr (Read Data)) >>= fun w__1 : TR_Result (mword 64) ExceptionType =>
+ (match w__1 with
+ | TR_Failure (e, _) => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ | TR_Address (paddr, _) =>
(match (width, 64) with
- | (BYTE, g__207) =>
- (mem_read Read addr 1 aq rl false) >>= fun w__3 : MemoryOpResult (mword (8 * 1)) =>
- (process_load rd vaddr w__3 is_unsigned)
+ | (BYTE, g__303) =>
+ (mem_read (Read Data) paddr 1 aq rl false) >>= fun w__2 : MemoryOpResult (mword (8 * 1)) =>
+ (process_load rd vaddr w__2 is_unsigned)
: M (Retired)
- | (HALF, g__208) =>
- (mem_read Read addr 2 aq rl false) >>= fun w__5 : MemoryOpResult (mword (8 * 2)) =>
- (process_load rd vaddr w__5 is_unsigned)
+ | (HALF, g__304) =>
+ (mem_read (Read Data) paddr 2 aq rl false) >>= fun w__4 : MemoryOpResult (mword (8 * 2)) =>
+ (process_load rd vaddr w__4 is_unsigned)
: M (Retired)
- | (WORD, g__209) =>
- (mem_read Read addr 4 aq rl false) >>= fun w__7 : MemoryOpResult (mword (8 * 4)) =>
- (process_load rd vaddr w__7 is_unsigned)
+ | (WORD, g__305) =>
+ (mem_read (Read Data) paddr 4 aq rl false) >>= fun w__6 : MemoryOpResult (mword (8 * 4)) =>
+ (process_load rd vaddr w__6 is_unsigned)
: M (Retired)
| (DOUBLE, l__0) =>
- (if sumbool_of_bool ((Z.eqb l__0 64)) then
- (mem_read Read addr 8 aq rl false) >>= fun w__9 : MemoryOpResult (mword (8 * 8)) =>
- (process_load rd vaddr w__9 is_unsigned)
+ (if sumbool_of_bool (Z.eqb l__0 64) then
+ (mem_read (Read Data) paddr 8 aq rl false) >>= fun w__8 : MemoryOpResult (mword (8 * 8)) =>
+ (process_load rd vaddr w__8 is_unsigned)
: M (Retired)
else
assert_exp' false "Pattern match failure at model/riscv_insts_base.sail 329:10 - 338:11" >>= fun _ =>
@@ -22578,10 +29284,8 @@ Definition execute_LOAD
end)
: M (Retired).
-Definition execute_ITYPE (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (op : iop)
-: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun rs1_val =>
+Definition execute_ITYPE (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (op : iop) : M (Retired) :=
+ (rX_bits rs1) >>= fun rs1_val =>
let immext : xlenbits := EXTS 64 imm in
let result : xlenbits :=
match op with
@@ -22592,227 +29296,634 @@ Definition execute_ITYPE (imm : mword 12) (rs1 : mword 5) (rd : mword 5) (op : i
| RISCV_ORI => or_vec rs1_val immext
| RISCV_XORI => xor_vec rs1_val immext
end in
- (wX (projT1 (regidx_to_regno rd)) result) >> returnm (RETIRE_SUCCESS : Retired).
+ (wX_bits rd result) >> returnm RETIRE_SUCCESS.
+
+Definition execute_ILLEGAL (s : mword 32) : M (Retired) :=
+ (handle_illegal tt) >> returnm RETIRE_FAIL.
+
+Definition execute_F_UN_TYPE_S (arg0 : mword 5) (arg1 : mword 5) (arg2 : f_un_op_S) : M (Retired) :=
+ let merge_var := (arg0, arg1, arg2) in
+ (match merge_var with
+ | (rs1, rd, FCLASS_S) =>
+ (rF_bits rs1) >>= fun w__0 : mword 64 =>
+ let rs1_val_S := nan_unbox w__0 in
+ let rd_val_10b : bits 10 :=
+ if f_is_neg_inf_S rs1_val_S then 'b"0000000001" : mword 10
+ else if f_is_neg_norm_S rs1_val_S then 'b"0000000010" : mword 10
+ else if f_is_neg_subnorm_S rs1_val_S then 'b"0000000100" : mword 10
+ else if f_is_neg_zero_S rs1_val_S then 'b"0000001000" : mword 10
+ else if f_is_pos_zero_S rs1_val_S then 'b"0000010000" : mword 10
+ else if f_is_pos_subnorm_S rs1_val_S then 'b"0000100000" : mword 10
+ else if f_is_pos_norm_S rs1_val_S then 'b"0001000000" : mword 10
+ else if f_is_pos_inf_S rs1_val_S then 'b"0010000000" : mword 10
+ else if f_is_SNaN_S rs1_val_S then 'b"0100000000" : mword 10
+ else if f_is_QNaN_S rs1_val_S then 'b"1000000000" : mword 10
+ else zeros_implicit 10 in
+ (wX_bits rd (EXTZ 64 rd_val_10b)) >> returnm RETIRE_SUCCESS
+ | (rs1, rd, FMV_X_W) =>
+ (rF_bits rs1) >>= fun w__1 : mword 64 =>
+ let rs1_val_S := subrange_vec_dec w__1 31 0 in
+ let rd_val_X : xlenbits := EXTS 64 rs1_val_S in
+ (wX_bits rd rd_val_X) >> returnm RETIRE_SUCCESS
+ | (rs1, rd, FMV_W_X) =>
+ (rX_bits rs1) >>= fun rs1_val_X =>
+ let rd_val_S := subrange_vec_dec rs1_val_X 31 0 in
+ (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ end)
+ : M (Retired).
-Definition execute_ILLEGAL (s : mword 32)
+Definition execute_F_UN_TYPE_D (arg0 : mword 5) (arg1 : mword 5) (arg2 : f_un_op_D) : M (Retired) :=
+ let merge_var := (arg0, arg1, arg2) in
+ (match merge_var with
+ | (rs1, rd, FCLASS_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ let rd_val_10b : bits 10 :=
+ if f_is_neg_inf_D rs1_val_D then 'b"0000000001" : mword 10
+ else if f_is_neg_norm_D rs1_val_D then 'b"0000000010" : mword 10
+ else if f_is_neg_subnorm_D rs1_val_D then 'b"0000000100" : mword 10
+ else if f_is_neg_zero_D rs1_val_D then 'b"0000001000" : mword 10
+ else if f_is_pos_zero_D rs1_val_D then 'b"0000010000" : mword 10
+ else if f_is_pos_subnorm_D rs1_val_D then 'b"0000100000" : mword 10
+ else if f_is_pos_norm_D rs1_val_D then 'b"0001000000" : mword 10
+ else if f_is_pos_inf_D rs1_val_D then 'b"0010000000" : mword 10
+ else if f_is_SNaN_D rs1_val_D then 'b"0100000000" : mword 10
+ else if f_is_QNaN_D rs1_val_D then 'b"1000000000" : mword 10
+ else zeros_implicit 10 in
+ (wX_bits rd (EXTZ 64 rd_val_10b)) >> returnm RETIRE_SUCCESS
+ | (rs1, rd, FMV_X_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ let rd_val_X := rs1_val_D in
+ (wX_bits rd rd_val_X) >> returnm RETIRE_SUCCESS
+ | (rs1, rd, FMV_D_X) =>
+ (rX_bits rs1) >>= fun rs1_val_X =>
+ let rd_val_D := rs1_val_X in
+ (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ end)
+ : M (Retired).
+
+Definition execute_F_UN_RM_TYPE_S
+(arg0 : mword 5) (arg1 : rounding_mode) (arg2 : mword 5) (arg3 : f_un_rm_op_S)
: M (Retired) :=
-
- (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired).
+ let merge_var := (arg0, arg1, arg2, arg3) in
+ (match merge_var with
+ | (rs1, rm, rd, FSQRT_S) =>
+ (rF_bits rs1) >>= fun w__0 : mword 64 =>
+ let rs1_val_S := nan_unbox w__0 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__1 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__1 in
+ (riscv_f32Sqrt rm_3b rs1_val_S) >>= fun '(fflags, rd_val_S) =>
+ (write_fflags fflags) >> (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_W_S) =>
+ (rF_bits rs1) >>= fun w__2 : mword 64 =>
+ let rs1_val_S := nan_unbox w__2 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__3 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__3 in
+ (riscv_f32ToI32 rm_3b rs1_val_S) >>= fun '(fflags, rd_val_W) =>
+ (write_fflags fflags) >> (wX_bits rd (EXTS 64 rd_val_W)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_WU_S) =>
+ (rF_bits rs1) >>= fun w__4 : mword 64 =>
+ let rs1_val_S := nan_unbox w__4 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__5 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__5 in
+ (riscv_f32ToUi32 rm_3b rs1_val_S) >>= fun '(fflags, rd_val_WU) =>
+ (write_fflags fflags) >> (wX_bits rd (EXTS 64 rd_val_WU)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_S_W) =>
+ (rX_bits rs1) >>= fun w__6 : mword 64 =>
+ let rs1_val_W := subrange_vec_dec w__6 31 0 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__7 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__7 in
+ (riscv_i32ToF32 rm_3b rs1_val_W) >>= fun '(fflags, rd_val_S) =>
+ (write_fflags fflags) >> (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_S_WU) =>
+ (rX_bits rs1) >>= fun w__8 : mword 64 =>
+ let rs1_val_WU := subrange_vec_dec w__8 31 0 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__9 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__9 in
+ (riscv_ui32ToF32 rm_3b rs1_val_WU) >>= fun '(fflags, rd_val_S) =>
+ (write_fflags fflags) >> (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_L_S) =>
+ (rF_bits rs1) >>= fun w__10 : mword 64 =>
+ let rs1_val_S := nan_unbox w__10 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__11 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__11 in
+ (riscv_f32ToI64 rm_3b rs1_val_S) >>= fun '(fflags, rd_val_L) =>
+ (write_fflags fflags) >> (wX_bits rd rd_val_L) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_LU_S) =>
+ (rF_bits rs1) >>= fun w__12 : mword 64 =>
+ let rs1_val_S := nan_unbox w__12 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__13 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__13 in
+ (riscv_f32ToUi64 rm_3b rs1_val_S) >>= fun '(fflags, rd_val_LU) =>
+ (write_fflags fflags) >> (wX_bits rd rd_val_LU) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_S_L) =>
+ (rX_bits rs1) >>= fun rs1_val_L =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__14 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__14 in
+ (riscv_i64ToF32 rm_3b rs1_val_L) >>= fun '(fflags, rd_val_S) =>
+ (write_fflags fflags) >> (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_S_LU) =>
+ (rX_bits rs1) >>= fun rs1_val_LU =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__15 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__15 in
+ (riscv_ui64ToF32 rm_3b rs1_val_LU) >>= fun '(fflags, rd_val_S) =>
+ (write_fflags fflags) >> (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ end)
+ : M (Retired).
+
+Definition execute_F_UN_RM_TYPE_D
+(arg0 : mword 5) (arg1 : rounding_mode) (arg2 : mword 5) (arg3 : f_un_rm_op_D)
+: M (Retired) :=
+ let merge_var := (arg0, arg1, arg2, arg3) in
+ (match merge_var with
+ | (rs1, rm, rd, FSQRT_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__0 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__0 in
+ (riscv_f64Sqrt rm_3b rs1_val_D) >>= fun '(fflags, rd_val_D) =>
+ (write_fflags fflags) >> (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_W_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__1 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__1 in
+ (riscv_f64ToI32 rm_3b rs1_val_D) >>= fun '(fflags, rd_val_W) =>
+ (write_fflags fflags) >> (wX_bits rd (EXTS 64 rd_val_W)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_WU_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__2 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__2 in
+ (riscv_f64ToUi32 rm_3b rs1_val_D) >>= fun '(fflags, rd_val_WU) =>
+ (write_fflags fflags) >> (wX_bits rd (EXTS 64 rd_val_WU)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_D_W) =>
+ (rX_bits rs1) >>= fun w__3 : mword 64 =>
+ let rs1_val_W := subrange_vec_dec w__3 31 0 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__4 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__4 in
+ (riscv_i32ToF64 rm_3b rs1_val_W) >>= fun '(fflags, rd_val_D) =>
+ (write_fflags fflags) >> (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_D_WU) =>
+ (rX_bits rs1) >>= fun w__5 : mword 64 =>
+ let rs1_val_WU := subrange_vec_dec w__5 31 0 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__6 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__6 in
+ (riscv_ui32ToF64 rm_3b rs1_val_WU) >>= fun '(fflags, rd_val_D) =>
+ (write_fflags fflags) >> (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_S_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__7 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__7 in
+ (riscv_f64ToF32 rm_3b rs1_val_D) >>= fun '(fflags, rd_val_S) =>
+ (write_fflags fflags) >> (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_D_S) =>
+ (rF_bits rs1) >>= fun w__8 : mword 64 =>
+ let rs1_val_S := nan_unbox w__8 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__9 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__9 in
+ (riscv_f32ToF64 rm_3b rs1_val_S) >>= fun '(fflags, rd_val_D) =>
+ (write_fflags fflags) >> (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_L_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__10 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__10 in
+ (riscv_f64ToI64 rm_3b rs1_val_D) >>= fun '(fflags, rd_val_L) =>
+ (write_fflags fflags) >> (wX_bits rd rd_val_L) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_LU_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__11 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__11 in
+ (riscv_f64ToUi64 rm_3b rs1_val_D) >>= fun '(fflags, rd_val_LU) =>
+ (write_fflags fflags) >> (wX_bits rd rd_val_LU) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_D_L) =>
+ (rX_bits rs1) >>= fun rs1_val_L =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__12 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__12 in
+ (riscv_i64ToF64 rm_3b rs1_val_L) >>= fun '(fflags, rd_val_D) =>
+ (write_fflags fflags) >> (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs1, rm, rd, FCVT_D_LU) =>
+ (rX_bits rs1) >>= fun rs1_val_LU =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__13 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__13 in
+ (riscv_ui64ToF64 rm_3b rs1_val_LU) >>= fun '(fflags, rd_val_D) =>
+ (write_fflags fflags) >> (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ end)
+ : M (Retired).
-Definition execute_FENCE_TSO (pred : mword 4) (succ : mword 4)
+Definition execute_F_MADD_TYPE_S
+(rs3 : mword 5) (rs2 : mword 5) (rs1 : mword 5) (rm : rounding_mode) (rd : mword 5)
+(op : f_madd_op_S)
: M (Retired) :=
-
+ (rF_bits rs1) >>= fun w__0 : mword 64 =>
+ let rs1_val_32b := nan_unbox w__0 in
+ (rF_bits rs2) >>= fun w__1 : mword 64 =>
+ let rs2_val_32b := nan_unbox w__1 in
+ (rF_bits rs3) >>= fun w__2 : mword 64 =>
+ let rs3_val_32b := nan_unbox w__2 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__3 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__3 in
+ (match op with
+ | FMADD_S =>
+ (riscv_f32MulAdd rm_3b rs1_val_32b rs2_val_32b rs3_val_32b) : M ((mword 5 * mword 32))
+ | FMSUB_S =>
+ (riscv_f32MulAdd rm_3b rs1_val_32b rs2_val_32b (negate_S rs3_val_32b))
+ : M ((mword 5 * mword 32))
+ | FNMSUB_S =>
+ (riscv_f32MulAdd rm_3b (negate_S rs1_val_32b) rs2_val_32b rs3_val_32b)
+ : M ((mword 5 * mword 32))
+ | FNMADD_S =>
+ (riscv_f32MulAdd rm_3b (negate_S rs1_val_32b) rs2_val_32b (negate_S rs3_val_32b))
+ : M ((mword 5 * mword 32))
+ end) >>= fun '((fflags, rd_val_32b)
+ : (bits 5 * bits 32)) =>
+ (write_fflags fflags) >> (wF_bits rd (nan_box rd_val_32b)) >> returnm RETIRE_SUCCESS.
+
+Definition execute_F_MADD_TYPE_D
+(rs3 : mword 5) (rs2 : mword 5) (rs1 : mword 5) (rm : rounding_mode) (rd : mword 5)
+(op : f_madd_op_D)
+: M (Retired) :=
+ (rF_bits rs1) >>= fun rs1_val_64b =>
+ (rF_bits rs2) >>= fun rs2_val_64b =>
+ (rF_bits rs3) >>= fun rs3_val_64b =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__0 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__0 in
+ (match op with
+ | FMADD_D =>
+ (riscv_f64MulAdd rm_3b rs1_val_64b rs2_val_64b rs3_val_64b) : M ((mword 5 * mword 64))
+ | FMSUB_D =>
+ (riscv_f64MulAdd rm_3b rs1_val_64b rs2_val_64b (negate_D rs3_val_64b))
+ : M ((mword 5 * mword 64))
+ | FNMSUB_D =>
+ (riscv_f64MulAdd rm_3b (negate_D rs1_val_64b) rs2_val_64b rs3_val_64b)
+ : M ((mword 5 * mword 64))
+ | FNMADD_D =>
+ (riscv_f64MulAdd rm_3b (negate_D rs1_val_64b) rs2_val_64b (negate_D rs3_val_64b))
+ : M ((mword 5 * mword 64))
+ end) >>= fun '((fflags, rd_val_64b)
+ : (bits 5 * bits 64)) =>
+ (write_fflags fflags) >> (wF_bits rd rd_val_64b) >> returnm RETIRE_SUCCESS.
+
+Definition execute_F_BIN_TYPE_S
+(arg0 : mword 5) (arg1 : mword 5) (arg2 : mword 5) (arg3 : f_bin_op_S)
+: M (Retired) :=
+ let merge_var := (arg0, arg1, arg2, arg3) in
+ (match merge_var with
+ | (rs2, rs1, rd, FSGNJ_S) =>
+ (rF_bits rs1) >>= fun w__0 : mword 64 =>
+ let rs1_val_S := nan_unbox w__0 in
+ (rF_bits rs2) >>= fun w__1 : mword 64 =>
+ let rs2_val_S := nan_unbox w__1 in
+ let '(s1, e1, m1) := fsplit_S rs1_val_S in
+ let '(s2, e2, m2) := fsplit_S rs2_val_S in
+ let rd_val_S := fmake_S s2 e1 m1 in
+ (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FSGNJN_S) =>
+ (rF_bits rs1) >>= fun w__2 : mword 64 =>
+ let rs1_val_S := nan_unbox w__2 in
+ (rF_bits rs2) >>= fun w__3 : mword 64 =>
+ let rs2_val_S := nan_unbox w__3 in
+ let '(s1, e1, m1) := fsplit_S rs1_val_S in
+ let '(s2, e2, m2) := fsplit_S rs2_val_S in
+ let rd_val_S := fmake_S (xor_vec ('b"1" : mword 1) s2) e1 m1 in
+ (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FSGNJX_S) =>
+ (rF_bits rs1) >>= fun w__4 : mword 64 =>
+ let rs1_val_S := nan_unbox w__4 in
+ (rF_bits rs2) >>= fun w__5 : mword 64 =>
+ let rs2_val_S := nan_unbox w__5 in
+ let '(s1, e1, m1) := fsplit_S rs1_val_S in
+ let '(s2, e2, m2) := fsplit_S rs2_val_S in
+ let rd_val_S := fmake_S (xor_vec s1 s2) e1 m1 in
+ (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FMIN_S) =>
+ (rF_bits rs1) >>= fun w__6 : mword 64 =>
+ let rs1_val_S := nan_unbox w__6 in
+ (rF_bits rs2) >>= fun w__7 : mword 64 =>
+ let rs2_val_S := nan_unbox w__7 in
+ let is_quiet := true in
+ let '(rs1_lt_rs2, fflags) := fle_S rs1_val_S rs2_val_S is_quiet in
+ let rd_val_S :=
+ if andb (f_is_NaN_S rs1_val_S) (f_is_NaN_S rs2_val_S) then canonical_NaN_S tt
+ else if f_is_NaN_S rs1_val_S then rs2_val_S
+ else if f_is_NaN_S rs2_val_S then rs1_val_S
+ else if andb (f_is_neg_zero_S rs1_val_S) (f_is_pos_zero_S rs2_val_S) then rs1_val_S
+ else if andb (f_is_neg_zero_S rs2_val_S) (f_is_pos_zero_S rs1_val_S) then rs2_val_S
+ else if sumbool_of_bool rs1_lt_rs2 then rs1_val_S
+ else rs2_val_S in
+ (accrue_fflags fflags) >> (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FMAX_S) =>
+ (rF_bits rs1) >>= fun w__8 : mword 64 =>
+ let rs1_val_S := nan_unbox w__8 in
+ (rF_bits rs2) >>= fun w__9 : mword 64 =>
+ let rs2_val_S := nan_unbox w__9 in
+ let is_quiet := true in
+ let '(rs2_lt_rs1, fflags) := fle_S rs2_val_S rs1_val_S is_quiet in
+ let rd_val_S :=
+ if andb (f_is_NaN_S rs1_val_S) (f_is_NaN_S rs2_val_S) then canonical_NaN_S tt
+ else if f_is_NaN_S rs1_val_S then rs2_val_S
+ else if f_is_NaN_S rs2_val_S then rs1_val_S
+ else if andb (f_is_neg_zero_S rs1_val_S) (f_is_pos_zero_S rs2_val_S) then rs2_val_S
+ else if andb (f_is_neg_zero_S rs2_val_S) (f_is_pos_zero_S rs1_val_S) then rs1_val_S
+ else if sumbool_of_bool rs2_lt_rs1 then rs1_val_S
+ else rs2_val_S in
+ (accrue_fflags fflags) >> (wF_bits rd (nan_box rd_val_S)) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FEQ_S) =>
+ (rF_bits rs1) >>= fun w__10 : mword 64 =>
+ let rs1_val_S := nan_unbox w__10 in
+ (rF_bits rs2) >>= fun w__11 : mword 64 =>
+ let rs2_val_S := nan_unbox w__11 in
+ (riscv_f32Eq rs1_val_S rs2_val_S) >>= fun '((fflags, rd_val)
+ : (bits_fflags * bits_WU)) =>
+ (write_fflags fflags) >> (wX_bits rd (EXTZ 64 rd_val)) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FLT_S) =>
+ (rF_bits rs1) >>= fun w__12 : mword 64 =>
+ let rs1_val_S := nan_unbox w__12 in
+ (rF_bits rs2) >>= fun w__13 : mword 64 =>
+ let rs2_val_S := nan_unbox w__13 in
+ (riscv_f32Lt rs1_val_S rs2_val_S) >>= fun '((fflags, rd_val)
+ : (bits_fflags * bits_WU)) =>
+ (write_fflags fflags) >> (wX_bits rd (EXTZ 64 rd_val)) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FLE_S) =>
+ (rF_bits rs1) >>= fun w__14 : mword 64 =>
+ let rs1_val_S := nan_unbox w__14 in
+ (rF_bits rs2) >>= fun w__15 : mword 64 =>
+ let rs2_val_S := nan_unbox w__15 in
+ (riscv_f32Le rs1_val_S rs2_val_S) >>= fun '((fflags, rd_val)
+ : (bits_fflags * bits_WU)) =>
+ (write_fflags fflags) >> (wX_bits rd (EXTZ 64 rd_val)) >> returnm RETIRE_SUCCESS
+ end)
+ : M (Retired).
+
+Definition execute_F_BIN_TYPE_D
+(arg0 : mword 5) (arg1 : mword 5) (arg2 : mword 5) (arg3 : f_bin_op_D)
+: M (Retired) :=
+ let merge_var := (arg0, arg1, arg2, arg3) in
+ (match merge_var with
+ | (rs2, rs1, rd, FSGNJ_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (rF_bits rs2) >>= fun rs2_val_D =>
+ let '(s1, e1, m1) := fsplit_D rs1_val_D in
+ let '(s2, e2, m2) := fsplit_D rs2_val_D in
+ let rd_val_D := fmake_D s2 e1 m1 in
+ (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FSGNJN_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (rF_bits rs2) >>= fun rs2_val_D =>
+ let '(s1, e1, m1) := fsplit_D rs1_val_D in
+ let '(s2, e2, m2) := fsplit_D rs2_val_D in
+ let rd_val_D := fmake_D (xor_vec ('b"1" : mword 1) s2) e1 m1 in
+ (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FSGNJX_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (rF_bits rs2) >>= fun rs2_val_D =>
+ let '(s1, e1, m1) := fsplit_D rs1_val_D in
+ let '(s2, e2, m2) := fsplit_D rs2_val_D in
+ let rd_val_D := fmake_D (xor_vec s1 s2) e1 m1 in
+ (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FMIN_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (rF_bits rs2) >>= fun rs2_val_D =>
+ let is_quiet := true in
+ let '(rs1_lt_rs2, fflags) := fle_D rs1_val_D rs2_val_D is_quiet in
+ let rd_val_D :=
+ if andb (f_is_NaN_D rs1_val_D) (f_is_NaN_D rs2_val_D) then canonical_NaN_D tt
+ else if f_is_NaN_D rs1_val_D then rs2_val_D
+ else if f_is_NaN_D rs2_val_D then rs1_val_D
+ else if andb (f_is_neg_zero_D rs1_val_D) (f_is_pos_zero_D rs2_val_D) then rs1_val_D
+ else if andb (f_is_neg_zero_D rs2_val_D) (f_is_pos_zero_D rs1_val_D) then rs2_val_D
+ else if sumbool_of_bool rs1_lt_rs2 then rs1_val_D
+ else rs2_val_D in
+ (accrue_fflags fflags) >> (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FMAX_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (rF_bits rs2) >>= fun rs2_val_D =>
+ let is_quiet := true in
+ let '(rs2_lt_rs1, fflags) := fle_D rs2_val_D rs1_val_D is_quiet in
+ let rd_val_D :=
+ if andb (f_is_NaN_D rs1_val_D) (f_is_NaN_D rs2_val_D) then canonical_NaN_D tt
+ else if f_is_NaN_D rs1_val_D then rs2_val_D
+ else if f_is_NaN_D rs2_val_D then rs1_val_D
+ else if andb (f_is_neg_zero_D rs1_val_D) (f_is_pos_zero_D rs2_val_D) then rs2_val_D
+ else if andb (f_is_neg_zero_D rs2_val_D) (f_is_pos_zero_D rs1_val_D) then rs1_val_D
+ else if sumbool_of_bool rs2_lt_rs1 then rs1_val_D
+ else rs2_val_D in
+ (accrue_fflags fflags) >> (wF_bits rd rd_val_D) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FEQ_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (rF_bits rs2) >>= fun rs2_val_D =>
+ (riscv_f64Eq rs1_val_D rs2_val_D) >>= fun '((fflags, rd_val)
+ : (bits_fflags * bits_LU)) =>
+ (write_fflags fflags) >> (wX_bits rd rd_val) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FLT_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (rF_bits rs2) >>= fun rs2_val_D =>
+ (riscv_f64Lt rs1_val_D rs2_val_D) >>= fun '((fflags, rd_val)
+ : (bits_fflags * bits_LU)) =>
+ (write_fflags fflags) >> (wX_bits rd rd_val) >> returnm RETIRE_SUCCESS
+ | (rs2, rs1, rd, FLE_D) =>
+ (rF_bits rs1) >>= fun rs1_val_D =>
+ (rF_bits rs2) >>= fun rs2_val_D =>
+ (riscv_f64Le rs1_val_D rs2_val_D) >>= fun '((fflags, rd_val)
+ : (bits_fflags * bits_LU)) =>
+ (write_fflags fflags) >> (wX_bits rd rd_val) >> returnm RETIRE_SUCCESS
+ end)
+ : M (Retired).
+
+Definition execute_F_BIN_RM_TYPE_S
+(rs2 : mword 5) (rs1 : mword 5) (rm : rounding_mode) (rd : mword 5) (op : f_bin_rm_op_S)
+: M (Retired) :=
+ (rF_bits rs1) >>= fun w__0 : mword 64 =>
+ let rs1_val_32b := nan_unbox w__0 in
+ (rF_bits rs2) >>= fun w__1 : mword 64 =>
+ let rs2_val_32b := nan_unbox w__1 in
+ (select_instr_or_fcsr_rm rm) >>= fun w__2 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__2 in
+ (match op with
+ | FADD_S => (riscv_f32Add rm_3b rs1_val_32b rs2_val_32b) : M ((mword 5 * mword 32))
+ | FSUB_S => (riscv_f32Sub rm_3b rs1_val_32b rs2_val_32b) : M ((mword 5 * mword 32))
+ | FMUL_S => (riscv_f32Mul rm_3b rs1_val_32b rs2_val_32b) : M ((mword 5 * mword 32))
+ | FDIV_S => (riscv_f32Div rm_3b rs1_val_32b rs2_val_32b) : M ((mword 5 * mword 32))
+ end) >>= fun '((fflags, rd_val_32b)
+ : (bits 5 * bits 32)) =>
+ (write_fflags fflags) >> (wF_bits rd (nan_box rd_val_32b)) >> returnm RETIRE_SUCCESS.
+
+Definition execute_F_BIN_RM_TYPE_D
+(rs2 : mword 5) (rs1 : mword 5) (rm : rounding_mode) (rd : mword 5) (op : f_bin_rm_op_D)
+: M (Retired) :=
+ (rF_bits rs1) >>= fun rs1_val_64b =>
+ (rF_bits rs2) >>= fun rs2_val_64b =>
+ (select_instr_or_fcsr_rm rm) >>= fun w__0 : rounding_mode =>
+ let rm_3b := encdec_rounding_mode_forwards w__0 in
+ (match op with
+ | FADD_D => (riscv_f64Add rm_3b rs1_val_64b rs2_val_64b) : M ((mword 5 * mword 64))
+ | FSUB_D => (riscv_f64Sub rm_3b rs1_val_64b rs2_val_64b) : M ((mword 5 * mword 64))
+ | FMUL_D => (riscv_f64Mul rm_3b rs1_val_64b rs2_val_64b) : M ((mword 5 * mword 64))
+ | FDIV_D => (riscv_f64Div rm_3b rs1_val_64b rs2_val_64b) : M ((mword 5 * mword 64))
+ end) >>= fun '((fflags, rd_val_64b)
+ : (bits 5 * bits 64)) =>
+ (write_fflags fflags) >> (wF_bits rd rd_val_64b) >> returnm RETIRE_SUCCESS.
+
+Definition execute_FENCE_TSO (pred : mword 4) (succ : mword 4) : M (Retired) :=
(match (pred, succ) with
- | (v__794, v__795) =>
- (if ((andb (eq_vec (subrange_vec_dec v__794 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__795 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_tso (tt)))
+ | (v__1330, v__1331) =>
+ (if andb (eq_vec (subrange_vec_dec v__1330 1 0) ('b"11" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1331 1 0) ('b"11" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_tso tt))
: M (unit)
else
- returnm ((if ((andb
- (eq_vec (subrange_vec_dec v__794 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__795 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1))))) then
- tt
- else
- let '_ := (print_endline "FIXME: unsupported fence") : unit in
- tt)
- : unit))
+ returnm (if andb (eq_vec (subrange_vec_dec v__1330 1 0) ('b"00" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1331 1 0) ('b"00" : mword (1 - 0 + 1))) then
+ tt
+ else
+ let '_ := (print_endline "FIXME: unsupported fence") : unit in
+ tt))
: M (unit)
end) >>
- returnm (RETIRE_SUCCESS
- : Retired).
+ returnm RETIRE_SUCCESS.
-Definition execute_FENCEI '(tt : unit) : Retired := RETIRE_SUCCESS.
+Definition execute_FENCEI '(tt : unit) : Retired := RETIRE_SUCCESS.
-Definition execute_FENCE (pred : mword 4) (succ : mword 4)
-: M (Retired) :=
-
+Definition execute_FENCE (pred : mword 4) (succ : mword 4) : M (Retired) :=
(match (pred, succ) with
- | (v__754, v__755) =>
- (if ((andb (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_rw_rw (tt)))
+ | (v__1290, v__1291) =>
+ (if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"11" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"11" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_rw_rw tt))
: M (unit)
- else if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_r_rw (tt)))
+ else if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"10" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"11" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_r_rw tt))
: M (unit)
- else if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_r_r (tt)))
+ else if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"10" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"10" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_r_r tt))
: M (unit)
- else if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_rw_w (tt)))
+ else if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"11" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_rw_w tt))
: M (unit)
- else if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_w_w (tt)))
+ else if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"01" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_w_w tt))
: M (unit)
- else if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_w_rw (tt)))
+ else if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"01" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"11" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_w_rw tt))
: M (unit)
- else if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_rw_r (tt)))
+ else if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"11" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"10" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_rw_r tt))
: M (unit)
- else if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_r_w (tt)))
+ else if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"10" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_r_w tt))
: M (unit)
- else if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0) (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0) (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))))
- then
- (barrier (Barrier_RISCV_w_r (tt)))
+ else if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"01" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"10" : mword (1 - 0 + 1))) then
+ (barrier (Barrier_RISCV_w_r tt))
: M (unit)
else
- returnm ((if ((andb
- (eq_vec (subrange_vec_dec v__754 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__755 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1))))) then
- tt
- else
- let '_ := (print_endline "FIXME: unsupported fence") : unit in
- tt)
- : unit))
+ returnm (if andb (eq_vec (subrange_vec_dec v__1290 1 0) ('b"00" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1291 1 0) ('b"00" : mword (1 - 0 + 1))) then
+ tt
+ else
+ let '_ := (print_endline "FIXME: unsupported fence") : unit in
+ tt))
: M (unit)
end) >>
- returnm (RETIRE_SUCCESS
- : Retired).
+ returnm RETIRE_SUCCESS.
-Definition execute_ECALL '(tt : unit)
-: M (Retired) :=
-
+Definition execute_ECALL '(tt : unit) : M (Retired) :=
read_reg cur_privilege_ref >>= fun w__0 : Privilege =>
let t : sync_exception :=
{| sync_exception_trap :=
(match w__0 with
- | User => E_U_EnvCall
- | Supervisor => E_S_EnvCall
- | Machine => E_M_EnvCall
+ | User => E_U_EnvCall tt
+ | Supervisor => E_S_EnvCall tt
+ | Machine => E_M_EnvCall tt
end);
sync_exception_excinfo := (None : option xlenbits);
sync_exception_ext := None |} in
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
((read_reg PC_ref) : M (mword 64)) >>= fun w__2 : mword 64 =>
- (exception_handler w__1 (CTL_TRAP (t)) w__2) >>= fun w__3 : mword 64 =>
- (set_next_pc w__3) >> returnm (RETIRE_FAIL : Retired).
+ (exception_handler w__1 (CTL_TRAP t) w__2) >>= fun w__3 : mword 64 =>
+ (set_next_pc w__3) >> returnm RETIRE_FAIL.
-Definition execute_EBREAK '(tt : unit)
-: M (Retired) :=
-
- ((read_reg PC_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
- (handle_mem_exception w__0 E_Breakpoint) >> returnm (RETIRE_FAIL : Retired).
+Definition execute_EBREAK '(tt : unit) : M (Retired) :=
+ (handle_exception (E_Breakpoint tt)) >> returnm RETIRE_FAIL.
-Definition execute_DIVW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
-: M (Retired) :=
-
+Definition execute_DIVW (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool) : M (Retired) :=
(haveMulDiv tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool (w__0) then
- (rX (projT1 (regidx_to_regno rs1))) >>= fun w__1 : mword 64 =>
+ (if sumbool_of_bool w__0 then
+ (rX_bits rs1) >>= fun w__1 : mword 64 =>
let rs1_val := subrange_vec_dec w__1 31 0 in
- (rX (projT1 (regidx_to_regno rs2))) >>= fun w__2 : mword 64 =>
+ (rX_bits rs2) >>= fun w__2 : mword 64 =>
let rs2_val := subrange_vec_dec w__2 31 0 in
- let rs1_int : Z :=
- if sumbool_of_bool (s) then projT1 (sint rs1_val)
- else projT1 (uint rs1_val) in
- let rs2_int : Z :=
- if sumbool_of_bool (s) then projT1 (sint rs2_val)
- else projT1 (uint rs2_val) in
- let q : Z := if sumbool_of_bool ((Z.eqb rs2_int 0)) then (-1) else Z.quot rs1_int rs2_int in
+ let rs1_int : Z := if sumbool_of_bool s then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
+ let rs2_int : Z := if sumbool_of_bool s then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
+ let q : Z := if sumbool_of_bool (Z.eqb rs2_int 0) then (-1) else Z.quot rs1_int rs2_int in
let q' : Z :=
- if sumbool_of_bool ((andb s (Z.gtb q (Z.sub (projT1 (pow2 31)) 1)))) then Z.sub 0 (pow 2 31)
+ if sumbool_of_bool (andb s (Z.gtb q (Z.sub (projT1 (pow2 31)) 1))) then Z.sub 0 (pow 2 31)
else q in
- (wX (projT1 (regidx_to_regno rd)) (EXTS 64 (to_bits 32 q'))) >>
- returnm (RETIRE_SUCCESS
- : Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ (wX_bits rd (EXTS 64 (to_bits 32 q'))) >> returnm RETIRE_SUCCESS
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired).
-Definition execute_DIV (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool)
-: M (Retired) :=
-
+Definition execute_DIV (rs2 : mword 5) (rs1 : mword 5) (rd : mword 5) (s : bool) : M (Retired) :=
(haveMulDiv tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool (w__0) then
- (rX (projT1 (regidx_to_regno rs1))) >>= fun rs1_val =>
- (rX (projT1 (regidx_to_regno rs2))) >>= fun rs2_val =>
- let rs1_int : Z :=
- if sumbool_of_bool (s) then projT1 (sint rs1_val)
- else projT1 (uint rs1_val) in
- let rs2_int : Z :=
- if sumbool_of_bool (s) then projT1 (sint rs2_val)
- else projT1 (uint rs2_val) in
- let q : Z := if sumbool_of_bool ((Z.eqb rs2_int 0)) then (-1) else Z.quot rs1_int rs2_int in
+ (if sumbool_of_bool w__0 then
+ (rX_bits rs1) >>= fun rs1_val =>
+ (rX_bits rs2) >>= fun rs2_val =>
+ let rs1_int : Z := if sumbool_of_bool s then projT1 (sint rs1_val) else projT1 (uint rs1_val) in
+ let rs2_int : Z := if sumbool_of_bool s then projT1 (sint rs2_val) else projT1 (uint rs2_val) in
+ let q : Z := if sumbool_of_bool (Z.eqb rs2_int 0) then (-1) else Z.quot rs1_int rs2_int in
let q' : Z :=
- if sumbool_of_bool ((andb s (Z.gtb q xlen_max_signed))) then xlen_min_signed
+ if sumbool_of_bool (andb s (Z.gtb q xlen_max_signed)) then xlen_min_signed
else q in
- (wX (projT1 (regidx_to_regno rd)) (to_bits 64 q')) >> returnm (RETIRE_SUCCESS : Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ (wX_bits rd (to_bits 64 q')) >> returnm RETIRE_SUCCESS
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired).
-Definition execute_C_NOP '(tt : unit) : Retired := RETIRE_SUCCESS.
+Definition execute_C_NOP '(tt : unit) : Retired := RETIRE_SUCCESS.
-Definition execute_C_ILLEGAL (s : mword 16)
-: M (Retired) :=
-
- (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired).
+Definition execute_C_ILLEGAL (s : mword 16) : M (Retired) :=
+ (handle_illegal tt) >> returnm RETIRE_FAIL.
-Definition execute_CSR (csr : mword 12) (rs1 : mword 5) (rd : mword 5) (is_imm : bool) (op : csrop)
+Definition execute_CSR (csr : mword 12) (rs1 : mword 5) (rd : mword 5) (is_imm : bool) (op : csrop)
: M (Retired) :=
-
- (if sumbool_of_bool (is_imm) then returnm ((EXTZ 64 rs1) : mword 64)
- else (rX (projT1 (regidx_to_regno rs1))) : M (mword 64)) >>= fun rs1_val : xlenbits =>
+ (if sumbool_of_bool is_imm then returnm (EXTZ 64 rs1)
+ else (rX_bits rs1) : M (mword 64)) >>= fun rs1_val : xlenbits =>
let isWrite : bool :=
match op with
| CSRRW => true
| _ =>
- if sumbool_of_bool (is_imm) then projT1 (neq_int (projT1 (uint rs1_val)) 0)
+ if sumbool_of_bool is_imm then projT1 (neq_int (projT1 (uint rs1_val)) 0)
else projT1 (neq_int (projT1 (uint rs1)) 0)
end in
read_reg cur_privilege_ref >>= fun w__1 : Privilege =>
(check_CSR csr w__1 isWrite) >>= fun w__2 : bool =>
- (if sumbool_of_bool ((negb w__2)) then (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired)
+ (if sumbool_of_bool (negb w__2) then (handle_illegal tt) >> returnm RETIRE_FAIL
else
- (readCSR csr) >>= fun csr_val =>
- (if sumbool_of_bool (isWrite) then
- let new_val : xlenbits :=
- match op with
- | CSRRW => rs1_val
- | CSRRS => or_vec csr_val rs1_val
- | CSRRC => and_vec csr_val (not_vec rs1_val)
- end in
- (writeCSR csr new_val)
- : M (unit)
- else returnm (tt : unit)) >>
- (wX (projT1 (regidx_to_regno rd)) csr_val) >> returnm (RETIRE_SUCCESS : Retired))
+ read_reg cur_privilege_ref >>= fun w__3 : Privilege =>
+ (if negb (ext_check_CSR csr w__3 isWrite) then
+ let '_ := (ext_check_CSR_fail tt) : unit in
+ returnm RETIRE_FAIL
+ else
+ (readCSR csr) >>= fun csr_val =>
+ (if sumbool_of_bool isWrite then
+ let new_val : xlenbits :=
+ match op with
+ | CSRRW => rs1_val
+ | CSRRS => or_vec csr_val rs1_val
+ | CSRRC => and_vec csr_val (not_vec rs1_val)
+ end in
+ (writeCSR csr new_val)
+ : M (unit)
+ else returnm tt) >>
+ (wX_bits rd csr_val) >> returnm RETIRE_SUCCESS)
+ : M (Retired))
: M (Retired).
-Definition execute_BTYPE (imm : mword 13) (rs2 : mword 5) (rs1 : mword 5) (op : bop)
-: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun rs1_val =>
- (rX (projT1 (regidx_to_regno rs2))) >>= fun rs2_val =>
+Definition execute_BTYPE (imm : mword 13) (rs2 : mword 5) (rs1 : mword 5) (op : bop) : M (Retired) :=
+ (rX_bits rs1) >>= fun rs1_val =>
+ (rX_bits rs2) >>= fun rs2_val =>
let taken : bool :=
match op with
| RISCV_BEQ => eq_vec rs1_val rs2_val
@@ -22824,68 +29935,75 @@ Definition execute_BTYPE (imm : mword 13) (rs2 : mword 5) (rs1 : mword 5) (op :
end in
((read_reg PC_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
let t : xlenbits := add_vec w__0 (EXTS 64 imm) in
- (if sumbool_of_bool (taken) then
+ (if sumbool_of_bool taken then
(match (ext_control_check_pc t) with
- | Ext_ControlAddr_Error (e) =>
+ | Ext_ControlAddr_Error e =>
let '_ := (ext_handle_control_check_error e) : unit in
- returnm (RETIRE_FAIL
- : Retired)
- | Ext_ControlAddr_OK (target) =>
+ returnm RETIRE_FAIL
+ | Ext_ControlAddr_OK target =>
(and_boolM ((bit_to_bool (access_vec_dec target 1)) : M (bool))
((haveRVC tt) >>= fun w__2 : bool => returnm ((negb w__2) : bool))) >>= fun w__3 : bool =>
- (if sumbool_of_bool (w__3) then
- (handle_mem_exception target E_Fetch_Addr_Align) >> returnm (RETIRE_FAIL : Retired)
- else (set_next_pc target) >> returnm (RETIRE_SUCCESS : Retired))
+ (if sumbool_of_bool w__3 then
+ (handle_mem_exception target (E_Fetch_Addr_Align tt)) >> returnm RETIRE_FAIL
+ else (set_next_pc target) >> returnm RETIRE_SUCCESS)
: M (Retired)
end)
: M (Retired)
- else returnm (RETIRE_SUCCESS : Retired))
+ else returnm RETIRE_SUCCESS)
: M (Retired).
Definition execute_AMO
(op : amoop) (aq : bool) (rl : bool) (rs2 : mword 5) (rs1 : mword 5) (width : word_width)
-(rd : mword 5)
+(rd : mword 5)
: M (Retired) :=
-
(haveAtomics tt) >>= fun w__0 : bool =>
- (if sumbool_of_bool (w__0) then
- (ext_data_get_addr rs1 (zeros_implicit 64) ReadWrite width) >>= fun w__1 : Ext_DataAddr_Check unit =>
+ (if sumbool_of_bool w__0 then
+ (ext_data_get_addr rs1 (zeros_implicit 64) (ReadWrite Data) width) >>= fun w__1 : Ext_DataAddr_Check unit =>
(match w__1 with
- | Ext_DataAddr_Error (e) =>
+ | Ext_DataAddr_Error e =>
let '_ := (ext_handle_data_check_error e) : unit in
- returnm (RETIRE_FAIL
- : Retired)
- | Ext_DataAddr_OK (vaddr) =>
- (translateAddr vaddr ReadWrite) >>= fun w__2 : TR_Result (mword 64) ExceptionType =>
+ returnm RETIRE_FAIL
+ | Ext_DataAddr_OK vaddr =>
+ (translateAddr vaddr (ReadWrite Data)) >>= fun w__2 : TR_Result (mword 64) ExceptionType =>
(match w__2 with
- | TR_Failure (e) => (handle_mem_exception vaddr e) >> returnm (RETIRE_FAIL : Retired)
- | TR_Address (addr) =>
+ | TR_Failure (e, _) => (handle_mem_exception vaddr e) >> returnm RETIRE_FAIL
+ | TR_Address (addr, _) =>
(match (width, 64) with
- | (WORD, g__220) =>
+ | (WORD, g__316) =>
(mem_write_ea addr 4 (andb aq rl) rl true) : M (MemoryOpResult unit)
| (DOUBLE, l__7) =>
- (if sumbool_of_bool ((Z.eqb l__7 64)) then
+ (if sumbool_of_bool (Z.eqb l__7 64) then
(mem_write_ea addr 8 (andb aq rl) rl true)
: M (MemoryOpResult unit)
else (internal_error "AMO expected WORD or DOUBLE") : M (MemoryOpResult unit))
: M (MemoryOpResult unit)
| _ => (internal_error "AMO expected WORD or DOUBLE") : M (MemoryOpResult unit)
end) >>= fun eares : MemoryOpResult unit =>
- (rX (projT1 (regidx_to_regno rs2))) >>= fun rs2_val : xlenbits =>
+ let is_unsigned : bool :=
+ match op with | AMOMINU => true | AMOMAXU => true | _ => false end in
+ (match width with
+ | WORD =>
+ (if sumbool_of_bool is_unsigned then
+ (rX_bits rs2) >>= fun w__8 : mword 64 =>
+ returnm (EXTZ 64 (subrange_vec_dec w__8 31 0))
+ else
+ (rX_bits rs2) >>= fun w__9 : mword 64 =>
+ returnm (EXTS 64 (subrange_vec_dec w__9 31 0)))
+ : M (mword 64)
+ | DOUBLE => (rX_bits rs2) : M (mword 64)
+ | _ => (internal_error "AMO expected WORD or DOUBLE") : M (mword 64)
+ end) >>= fun rs2_val : xlenbits =>
(match eares with
- | MemException (e) =>
- (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
- | MemValue (_) =>
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
+ | MemValue _ =>
(match (width, 64) with
- | (WORD, g__219) =>
- (mem_read ReadWrite addr 4 aq (andb aq rl) true) >>= fun w__8 : MemoryOpResult (mword (8 * 4)) =>
- returnm ((extend_value false w__8)
- : MemoryOpResult (mword 64))
+ | (WORD, g__315) =>
+ (mem_read (ReadWrite Data) addr 4 aq (andb aq rl) true) >>= fun w__13 : MemoryOpResult (mword (8 * 4)) =>
+ returnm (extend_value is_unsigned w__13)
| (DOUBLE, l__6) =>
- (if sumbool_of_bool ((Z.eqb l__6 64)) then
- (mem_read ReadWrite addr 8 aq (andb aq rl) true) >>= fun w__9 : MemoryOpResult (mword (8 * 8)) =>
- returnm ((extend_value false w__9)
- : MemoryOpResult (mword 64))
+ (if sumbool_of_bool (Z.eqb l__6 64) then
+ (mem_read (ReadWrite Data) addr 8 aq (andb aq rl) true) >>= fun w__14 : MemoryOpResult (mword (8 * 8)) =>
+ returnm (extend_value is_unsigned w__14)
else
(internal_error "AMO expected WORD or DOUBLE")
: M (MemoryOpResult (mword 64)))
@@ -22893,11 +30011,10 @@ Definition execute_AMO
| _ =>
(internal_error "AMO expected WORD or DOUBLE")
: M (MemoryOpResult (mword 64))
- end) >>= fun rval : MemoryOpResult xlenbits =>
- (match rval with
- | MemException (e) =>
- (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
- | MemValue (loaded) =>
+ end) >>= fun mval : MemoryOpResult xlenbits =>
+ (match mval with
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
+ | MemValue loaded =>
let result : xlenbits :=
match op with
| AMOSWAP => rs2_val
@@ -22906,25 +30023,34 @@ Definition execute_AMO
| AMOAND => and_vec rs2_val loaded
| AMOOR => or_vec rs2_val loaded
| AMOMIN =>
- to_bits 64 (Z.min (projT1 (sint rs2_val)) (projT1 (sint loaded)))
+ to_bits 64
+ (projT1
+ (min_atom (projT1 (sint rs2_val)) (projT1 (sint loaded))))
| AMOMAX =>
- to_bits 64 (Z.max (projT1 (sint rs2_val)) (projT1 (sint loaded)))
+ to_bits 64
+ (projT1
+ (max_atom (projT1 (sint rs2_val)) (projT1 (sint loaded))))
| AMOMINU =>
to_bits 64
(projT1
- (min_nat (projT1 (uint rs2_val)) (projT1 (uint loaded))))
+ (min_atom (projT1 (uint rs2_val)) (projT1 (uint loaded))))
| AMOMAXU =>
to_bits 64
(projT1
- (max_nat (projT1 (uint rs2_val)) (projT1 (uint loaded))))
+ (max_atom (projT1 (uint rs2_val)) (projT1 (uint loaded))))
end in
+ (match width with
+ | WORD => returnm (EXTS 64 (subrange_vec_dec loaded 31 0))
+ | DOUBLE => returnm loaded
+ | _ => (internal_error "AMO expected WORD or DOUBLE") : M (mword 64)
+ end) >>= fun rval : xlenbits =>
(match (width, 64) with
- | (WORD, g__218) =>
+ | (WORD, g__314) =>
(mem_write_value addr 4 (subrange_vec_dec result 31 0) (andb aq rl) rl
true)
: M (MemoryOpResult bool)
| (DOUBLE, l__5) =>
- (if sumbool_of_bool ((Z.eqb l__5 64)) then
+ (if sumbool_of_bool (Z.eqb l__5 64) then
(mem_write_value addr 8 result (andb aq rl) rl true)
: M (MemoryOpResult bool)
else
@@ -22935,14 +30061,10 @@ Definition execute_AMO
(internal_error "AMO expected WORD or DOUBLE") : M (MemoryOpResult bool)
end) >>= fun wval : MemoryOpResult bool =>
(match wval with
- | MemValue (true) =>
- (wX (projT1 (regidx_to_regno rd)) loaded) >>
- returnm (RETIRE_SUCCESS
- : Retired)
- | MemValue (false) =>
+ | MemValue true => (wX_bits rd rval) >> returnm RETIRE_SUCCESS
+ | MemValue false =>
(internal_error "AMO got false from mem_write_value") : M (Retired)
- | MemException (e) =>
- (handle_mem_exception addr e) >> returnm (RETIRE_FAIL : Retired)
+ | MemException e => (handle_mem_exception addr e) >> returnm RETIRE_FAIL
end)
: M (Retired)
end)
@@ -22953,6752 +30075,12853 @@ Definition execute_AMO
: M (Retired)
end)
: M (Retired)
- else (handle_illegal tt) >> returnm (RETIRE_FAIL : Retired))
+ else (handle_illegal tt) >> returnm RETIRE_FAIL)
: M (Retired).
-Definition execute_ADDIW (imm : mword 12) (rs1 : mword 5) (rd : mword 5)
-: M (Retired) :=
-
- (rX (projT1 (regidx_to_regno rs1))) >>= fun w__0 : mword 64 =>
+Definition execute_ADDIW (imm : mword 12) (rs1 : mword 5) (rd : mword 5) : M (Retired) :=
+ (rX_bits rs1) >>= fun w__0 : mword 64 =>
let result : xlenbits := add_vec (EXTS 64 imm) w__0 in
- (wX (projT1 (regidx_to_regno rd)) (EXTS 64 (subrange_vec_dec result 31 0))) >>
- returnm (RETIRE_SUCCESS
- : Retired).
+ (wX_bits rd (EXTS 64 (subrange_vec_dec result 31 0))) >> returnm RETIRE_SUCCESS.
-Definition compressed_measure (instr : ast)
-: Z :=
-
+Definition compressed_measure (instr : ast) : Z :=
match instr with
- | C_ADDI4SPN ((rdc, nzimm)) => 1
- | C_LW ((uimm, rsc, rdc)) => 1
- | C_LD ((uimm, rsc, rdc)) => 1
- | C_SW ((uimm, rsc1, rsc2)) => 1
- | C_SD ((uimm, rsc1, rsc2)) => 1
- | C_ADDI ((nzi, rsd)) => 1
- | C_JAL (imm) => 1
- | C_LI ((imm, rd)) => 1
- | C_ADDI16SP (imm) => 1
- | C_LUI ((imm, rd)) => 1
- | C_SRLI ((shamt, rsd)) => 1
- | C_SRAI ((shamt, rsd)) => 1
- | C_ANDI ((imm, rsd)) => 1
- | C_SUB ((rsd, rs2)) => 1
- | C_XOR ((rsd, rs2)) => 1
- | C_OR ((rsd, rs2)) => 1
- | C_AND ((rsd, rs2)) => 1
- | C_SUBW ((rsd, rs2)) => 1
- | C_ADDW ((rsd, rs2)) => 1
- | C_J (imm) => 1
- | C_BEQZ ((imm, rs)) => 1
- | C_BNEZ ((imm, rs)) => 1
- | C_SLLI ((shamt, rsd)) => 1
- | C_LWSP ((uimm, rd)) => 1
- | C_LDSP ((uimm, rd)) => 1
- | C_SWSP ((uimm, rs2)) => 1
- | C_SDSP ((uimm, rs2)) => 1
- | C_JR (rs1) => 1
- | C_JALR (rs1) => 1
- | C_MV ((rd, rs2)) => 1
- | C_EBREAK (tt') => 1
- | C_ADD ((rsd, rs2)) => 1
+ | C_ADDI4SPN (rdc, nzimm) => 1
+ | C_LW (uimm, rsc, rdc) => 1
+ | C_LD (uimm, rsc, rdc) => 1
+ | C_SW (uimm, rsc1, rsc2) => 1
+ | C_SD (uimm, rsc1, rsc2) => 1
+ | C_ADDI (nzi, rsd) => 1
+ | C_JAL imm => 1
+ | C_LI (imm, rd) => 1
+ | C_ADDI16SP imm => 1
+ | C_LUI (imm, rd) => 1
+ | C_SRLI (shamt, rsd) => 1
+ | C_SRAI (shamt, rsd) => 1
+ | C_ANDI (imm, rsd) => 1
+ | C_SUB (rsd, rs2) => 1
+ | C_XOR (rsd, rs2) => 1
+ | C_OR (rsd, rs2) => 1
+ | C_AND (rsd, rs2) => 1
+ | C_SUBW (rsd, rs2) => 1
+ | C_ADDW (rsd, rs2) => 1
+ | C_J imm => 1
+ | C_BEQZ (imm, rs) => 1
+ | C_BNEZ (imm, rs) => 1
+ | C_SLLI (shamt, rsd) => 1
+ | C_LWSP (uimm, rd) => 1
+ | C_LDSP (uimm, rd) => 1
+ | C_SWSP (uimm, rs2) => 1
+ | C_SDSP (uimm, rs2) => 1
+ | C_JR rs1 => 1
+ | C_JALR rs1 => 1
+ | C_MV (rd, rs2) => 1
+ | C_EBREAK tt' => 1
+ | C_ADD (rsd, rs2) => 1
| _ => 0
end.
-Fixpoint _rec_execute (merge_var : ast) (_reclimit : Z) (_acc : Acc (Zwf 0) _reclimit)
-{struct _acc} : M (Retired) :=
-
+Fixpoint _rec_execute (merge_var : ast) (_reclimit : Z) (_acc : Acc (Zwf 0) _reclimit)
+{struct _acc} : M (Retired).
+exact (
assert_exp' (Z.geb _reclimit 0) "recursion limit reached" >>= fun _ =>
(match merge_var with
- | C_ADDI4SPN ((rdc, nzimm)) =>
- let imm : bits 12 :=
- concat_vec (vec_of_bits [B0;B0] : mword 2)
- (concat_vec nzimm (vec_of_bits [B0;B0] : mword 2)) in
+ | C_ADDI4SPN (rdc, nzimm) =>
+ let imm : bits 12 := concat_vec ('b"00" : mword 2) (concat_vec nzimm ('b"00" : mword 2)) in
let rd := creg2reg_idx rdc in
- (_rec_execute (ITYPE ((imm, sp, rd, RISCV_ADDI))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ (_rec_execute (ITYPE (imm, sp, rd, RISCV_ADDI)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_LW ((uimm, rsc, rdc)) =>
- let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0] : mword 2)) in
+ | C_LW (uimm, rsc, rdc) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"00" : mword 2)) in
let rd := creg2reg_idx rdc in
let rs := creg2reg_idx rsc in
- (_rec_execute (LOAD ((imm, rs, rd, false, WORD, false, false))) (Z.sub _reclimit 1)
+ (_rec_execute (LOAD (imm, rs, rd, false, WORD, false, false)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_LD ((uimm, rsc, rdc)) =>
- let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0;B0] : mword 3)) in
+ | C_LD (uimm, rsc, rdc) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"000" : mword 3)) in
let rd := creg2reg_idx rdc in
let rs := creg2reg_idx rsc in
- (_rec_execute (LOAD ((imm, rs, rd, false, DOUBLE, false, false))) (Z.sub _reclimit 1)
+ (_rec_execute (LOAD (imm, rs, rd, false, DOUBLE, false, false)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_SW ((uimm, rsc1, rsc2)) =>
- let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0] : mword 2)) in
+ | C_SW (uimm, rsc1, rsc2) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"00" : mword 2)) in
let rs1 := creg2reg_idx rsc1 in
let rs2 := creg2reg_idx rsc2 in
- (_rec_execute (STORE ((imm, rs2, rs1, WORD, false, false))) (Z.sub _reclimit 1)
+ (_rec_execute (STORE (imm, rs2, rs1, WORD, false, false)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_SD ((uimm, rsc1, rsc2)) =>
- let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0;B0] : mword 3)) in
+ | C_SD (uimm, rsc1, rsc2) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"000" : mword 3)) in
let rs1 := creg2reg_idx rsc1 in
let rs2 := creg2reg_idx rsc2 in
- (_rec_execute (STORE ((imm, rs2, rs1, DOUBLE, false, false))) (Z.sub _reclimit 1)
+ (_rec_execute (STORE (imm, rs2, rs1, DOUBLE, false, false)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_ADDI ((nzi, rsd)) =>
+ | C_ADDI (nzi, rsd) =>
let imm : bits 12 := EXTS 12 nzi in
- (_rec_execute (ITYPE ((imm, rsd, rsd, RISCV_ADDI))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ (_rec_execute (ITYPE (imm, rsd, rsd, RISCV_ADDI)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_JAL (imm) =>
- (_rec_execute (RISCV_JAL ((EXTS 21 (concat_vec imm (vec_of_bits [B0] : mword 1)), ra)))
+ | C_JAL imm =>
+ (_rec_execute (RISCV_JAL (EXTS 21 (concat_vec imm ('b"0" : mword 1)), ra))
(Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_ADDIW ((imm, rsd)) =>
- (_rec_execute (ADDIW ((EXTS 12 imm, rsd, rsd))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ | C_ADDIW (imm, rsd) =>
+ (_rec_execute (ADDIW (EXTS 12 imm, rsd, rsd)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_LI ((imm, rd)) =>
+ | C_LI (imm, rd) =>
let imm : bits 12 := EXTS 12 imm in
- (_rec_execute (ITYPE ((imm, zreg, rd, RISCV_ADDI))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ (_rec_execute (ITYPE (imm, zreg, rd, RISCV_ADDI)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_ADDI16SP (imm) =>
- let imm : bits 12 := EXTS 12 (concat_vec imm (vec_of_bits [B0;B0;B0;B0] : mword 4)) in
- (_rec_execute (ITYPE ((imm, sp, sp, RISCV_ADDI))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ | C_ADDI16SP imm =>
+ let imm : bits 12 := EXTS 12 (concat_vec imm (Ox"0" : mword 4)) in
+ (_rec_execute (ITYPE (imm, sp, sp, RISCV_ADDI)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_LUI ((imm, rd)) =>
+ | C_LUI (imm, rd) =>
let res : bits 20 := EXTS 20 imm in
- (_rec_execute (UTYPE ((res, rd, RISCV_LUI))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ (_rec_execute (UTYPE (res, rd, RISCV_LUI)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_SRLI ((shamt, rsd)) =>
+ | C_SRLI (shamt, rsd) =>
let rsd := creg2reg_idx rsd in
- (_rec_execute (SHIFTIOP ((shamt, rsd, rsd, RISCV_SRLI))) (Z.sub _reclimit 1)
+ (_rec_execute (SHIFTIOP (shamt, rsd, rsd, RISCV_SRLI)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_SRAI ((shamt, rsd)) =>
+ | C_SRAI (shamt, rsd) =>
let rsd := creg2reg_idx rsd in
- (_rec_execute (SHIFTIOP ((shamt, rsd, rsd, RISCV_SRAI))) (Z.sub _reclimit 1)
+ (_rec_execute (SHIFTIOP (shamt, rsd, rsd, RISCV_SRAI)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_ANDI ((imm, rsd)) =>
+ | C_ANDI (imm, rsd) =>
let rsd := creg2reg_idx rsd in
- (_rec_execute (ITYPE ((EXTS 12 imm, rsd, rsd, RISCV_ANDI))) (Z.sub _reclimit 1)
+ (_rec_execute (ITYPE (EXTS 12 imm, rsd, rsd, RISCV_ANDI)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_SUB ((rsd, rs2)) =>
+ | C_SUB (rsd, rs2) =>
let rsd := creg2reg_idx rsd in
let rs2 := creg2reg_idx rs2 in
- (_rec_execute (RTYPE ((rs2, rsd, rsd, RISCV_SUB))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ (_rec_execute (RTYPE (rs2, rsd, rsd, RISCV_SUB)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_XOR ((rsd, rs2)) =>
+ | C_XOR (rsd, rs2) =>
let rsd := creg2reg_idx rsd in
let rs2 := creg2reg_idx rs2 in
- (_rec_execute (RTYPE ((rs2, rsd, rsd, RISCV_XOR))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ (_rec_execute (RTYPE (rs2, rsd, rsd, RISCV_XOR)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_OR ((rsd, rs2)) =>
+ | C_OR (rsd, rs2) =>
let rsd := creg2reg_idx rsd in
let rs2 := creg2reg_idx rs2 in
- (_rec_execute (RTYPE ((rs2, rsd, rsd, RISCV_OR))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ (_rec_execute (RTYPE (rs2, rsd, rsd, RISCV_OR)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_AND ((rsd, rs2)) =>
+ | C_AND (rsd, rs2) =>
let rsd := creg2reg_idx rsd in
let rs2 := creg2reg_idx rs2 in
- (_rec_execute (RTYPE ((rs2, rsd, rsd, RISCV_AND))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ (_rec_execute (RTYPE (rs2, rsd, rsd, RISCV_AND)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_SUBW ((rsd, rs2)) =>
+ | C_SUBW (rsd, rs2) =>
let rsd := creg2reg_idx rsd in
let rs2 := creg2reg_idx rs2 in
- (_rec_execute (RTYPEW ((rs2, rsd, rsd, RISCV_SUBW))) (Z.sub _reclimit 1)
- (_limit_reduces _acc))
+ (_rec_execute (RTYPEW (rs2, rsd, rsd, RISCV_SUBW)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_ADDW ((rsd, rs2)) =>
+ | C_ADDW (rsd, rs2) =>
let rsd := creg2reg_idx rsd in
let rs2 := creg2reg_idx rs2 in
- (_rec_execute (RTYPEW ((rs2, rsd, rsd, RISCV_ADDW))) (Z.sub _reclimit 1)
- (_limit_reduces _acc))
+ (_rec_execute (RTYPEW (rs2, rsd, rsd, RISCV_ADDW)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_J (imm) =>
- (_rec_execute (RISCV_JAL ((EXTS 21 (concat_vec imm (vec_of_bits [B0] : mword 1)), zreg)))
+ | C_J imm =>
+ (_rec_execute (RISCV_JAL (EXTS 21 (concat_vec imm ('b"0" : mword 1)), zreg))
(Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_BEQZ ((imm, rs)) =>
+ | C_BEQZ (imm, rs) =>
(_rec_execute
- (BTYPE
- ((EXTS 13 (concat_vec imm (vec_of_bits [B0] : mword 1)), zreg, creg2reg_idx rs, RISCV_BEQ)))
+ (BTYPE (EXTS 13 (concat_vec imm ('b"0" : mword 1)), zreg, creg2reg_idx rs, RISCV_BEQ))
(Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_BNEZ ((imm, rs)) =>
+ | C_BNEZ (imm, rs) =>
(_rec_execute
- (BTYPE
- ((EXTS 13 (concat_vec imm (vec_of_bits [B0] : mword 1)), zreg, creg2reg_idx rs, RISCV_BNE)))
+ (BTYPE (EXTS 13 (concat_vec imm ('b"0" : mword 1)), zreg, creg2reg_idx rs, RISCV_BNE))
(Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_SLLI ((shamt, rsd)) =>
- (_rec_execute (SHIFTIOP ((shamt, rsd, rsd, RISCV_SLLI))) (Z.sub _reclimit 1)
+ | C_SLLI (shamt, rsd) =>
+ (_rec_execute (SHIFTIOP (shamt, rsd, rsd, RISCV_SLLI)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_LWSP ((uimm, rd)) =>
- let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0] : mword 2)) in
- (_rec_execute (LOAD ((imm, sp, rd, false, WORD, false, false))) (Z.sub _reclimit 1)
+ | C_LWSP (uimm, rd) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"00" : mword 2)) in
+ (_rec_execute (LOAD (imm, sp, rd, false, WORD, false, false)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_LDSP ((uimm, rd)) =>
- let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0;B0] : mword 3)) in
- (_rec_execute (LOAD ((imm, sp, rd, false, DOUBLE, false, false))) (Z.sub _reclimit 1)
+ | C_LDSP (uimm, rd) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"000" : mword 3)) in
+ (_rec_execute (LOAD (imm, sp, rd, false, DOUBLE, false, false)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_SWSP ((uimm, rs2)) =>
- let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0] : mword 2)) in
- (_rec_execute (STORE ((imm, rs2, sp, WORD, false, false))) (Z.sub _reclimit 1)
+ | C_SWSP (uimm, rs2) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"00" : mword 2)) in
+ (_rec_execute (STORE (imm, rs2, sp, WORD, false, false)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_SDSP ((uimm, rs2)) =>
- let imm : bits 12 := EXTZ 12 (concat_vec uimm (vec_of_bits [B0;B0;B0] : mword 3)) in
- (_rec_execute (STORE ((imm, rs2, sp, DOUBLE, false, false))) (Z.sub _reclimit 1)
+ | C_SDSP (uimm, rs2) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"000" : mword 3)) in
+ (_rec_execute (STORE (imm, rs2, sp, DOUBLE, false, false)) (Z.sub _reclimit 1)
(_limit_reduces _acc))
: M (Retired)
- | C_JR (rs1) =>
- (_rec_execute (RISCV_JALR ((EXTZ 12 (vec_of_bits [B0] : mword 1), rs1, zreg)))
- (Z.sub _reclimit 1) (_limit_reduces _acc))
+ | C_JR rs1 =>
+ (_rec_execute (RISCV_JALR (EXTZ 12 ('b"0" : mword 1), rs1, zreg)) (Z.sub _reclimit 1)
+ (_limit_reduces _acc))
: M (Retired)
- | C_JALR (rs1) =>
- (_rec_execute (RISCV_JALR ((EXTZ 12 (vec_of_bits [B0] : mword 1), rs1, ra)))
- (Z.sub _reclimit 1) (_limit_reduces _acc))
+ | C_JALR rs1 =>
+ (_rec_execute (RISCV_JALR (EXTZ 12 ('b"0" : mword 1), rs1, ra)) (Z.sub _reclimit 1)
+ (_limit_reduces _acc))
+ : M (Retired)
+ | C_MV (rd, rs2) =>
+ (_rec_execute (RTYPE (rs2, zreg, rd, RISCV_ADD)) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ : M (Retired)
+ | C_EBREAK tt =>
+ (_rec_execute (EBREAK tt) (Z.sub _reclimit 1) (_limit_reduces _acc)) : M (Retired)
+ | C_ADD (rsd, rs2) =>
+ (_rec_execute (RTYPE (rs2, rsd, rsd, RISCV_ADD)) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ : M (Retired)
+ | C_FLWSP (imm, rd) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec imm ('b"00" : mword 2)) in
+ (_rec_execute (LOAD_FP (imm, sp, rd, WORD)) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ : M (Retired)
+ | C_FSWSP (uimm, rs2) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"00" : mword 2)) in
+ (_rec_execute (STORE_FP (imm, rs2, sp, WORD)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_MV ((rd, rs2)) =>
- (_rec_execute (RTYPE ((rs2, zreg, rd, RISCV_ADD))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ | C_FLW (uimm, rsc, rdc) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"00" : mword 2)) in
+ let rd := creg2reg_idx rdc in
+ let rs := creg2reg_idx rsc in
+ (_rec_execute (LOAD_FP (imm, rs, rd, WORD)) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ : M (Retired)
+ | C_FSW (uimm, rsc1, rsc2) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"00" : mword 2)) in
+ let rs1 := creg2reg_idx rsc1 in
+ let rs2 := creg2reg_idx rsc2 in
+ (_rec_execute (STORE_FP (imm, rs2, rs1, WORD)) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ : M (Retired)
+ | C_FLDSP (uimm, rd) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"000" : mword 3)) in
+ (_rec_execute (LOAD_FP (imm, sp, rd, DOUBLE)) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ : M (Retired)
+ | C_FSDSP (uimm, rs2) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"000" : mword 3)) in
+ (_rec_execute (STORE_FP (imm, rs2, sp, DOUBLE)) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ : M (Retired)
+ | C_FLD (uimm, rsc, rdc) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"000" : mword 3)) in
+ let rd := creg2reg_idx rdc in
+ let rs := creg2reg_idx rsc in
+ (_rec_execute (LOAD_FP (imm, rs, rd, DOUBLE)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | C_EBREAK (tt) =>
- (_rec_execute (EBREAK (tt)) (Z.sub _reclimit 1) (_limit_reduces _acc)) : M (Retired)
- | C_ADD ((rsd, rs2)) =>
- (_rec_execute (RTYPE ((rs2, rsd, rsd, RISCV_ADD))) (Z.sub _reclimit 1) (_limit_reduces _acc))
+ | C_FSD (uimm, rsc1, rsc2) =>
+ let imm : bits 12 := EXTZ 12 (concat_vec uimm ('b"000" : mword 3)) in
+ let rs1 := creg2reg_idx rsc1 in
+ let rs2 := creg2reg_idx rsc2 in
+ (_rec_execute (STORE_FP (imm, rs2, rs1, DOUBLE)) (Z.sub _reclimit 1) (_limit_reduces _acc))
: M (Retired)
- | UTYPE ((imm, rd, op)) => (execute_UTYPE imm rd op) : M (Retired)
- | RISCV_JAL ((imm, rd)) => (execute_RISCV_JAL imm rd) : M (Retired)
- | BTYPE ((imm, rs2, rs1, op)) => (execute_BTYPE imm rs2 rs1 op) : M (Retired)
- | ITYPE ((imm, rs1, rd, op)) => (execute_ITYPE imm rs1 rd op) : M (Retired)
- | SHIFTIOP ((shamt, rs1, rd, op)) => (execute_SHIFTIOP shamt rs1 rd op) : M (Retired)
- | RTYPE ((rs2, rs1, rd, op)) => (execute_RTYPE rs2 rs1 rd op) : M (Retired)
- | LOAD ((imm, rs1, rd, is_unsigned, width, aq, rl)) =>
+ | UTYPE (imm, rd, op) => (execute_UTYPE imm rd op) : M (Retired)
+ | RISCV_JAL (imm, rd) => (execute_RISCV_JAL imm rd) : M (Retired)
+ | BTYPE (imm, rs2, rs1, op) => (execute_BTYPE imm rs2 rs1 op) : M (Retired)
+ | ITYPE (imm, rs1, rd, op) => (execute_ITYPE imm rs1 rd op) : M (Retired)
+ | SHIFTIOP (shamt, rs1, rd, op) => (execute_SHIFTIOP shamt rs1 rd op) : M (Retired)
+ | RTYPE (rs2, rs1, rd, op) => (execute_RTYPE rs2 rs1 rd op) : M (Retired)
+ | LOAD (imm, rs1, rd, is_unsigned, width, aq, rl) =>
(execute_LOAD imm rs1 rd is_unsigned width aq rl) : M (Retired)
- | STORE ((imm, rs2, rs1, width, aq, rl)) =>
- (execute_STORE imm rs2 rs1 width aq rl) : M (Retired)
- | ADDIW ((imm, rs1, rd)) => (execute_ADDIW imm rs1 rd) : M (Retired)
- | SHIFTW ((shamt, rs1, rd, op)) => (execute_SHIFTW shamt rs1 rd op) : M (Retired)
- | RTYPEW ((rs2, rs1, rd, op)) => (execute_RTYPEW rs2 rs1 rd op) : M (Retired)
- | SHIFTIWOP ((shamt, rs1, rd, op)) => (execute_SHIFTIWOP shamt rs1 rd op) : M (Retired)
- | FENCE ((pred, succ)) => (execute_FENCE pred succ) : M (Retired)
- | FENCE_TSO ((pred, succ)) => (execute_FENCE_TSO pred succ) : M (Retired)
- | FENCEI (arg0) => returnm ((execute_FENCEI arg0) : Retired)
- | ECALL (arg0) => (execute_ECALL arg0) : M (Retired)
- | MRET (arg0) => (execute_MRET arg0) : M (Retired)
- | SRET (arg0) => (execute_SRET arg0) : M (Retired)
- | EBREAK (arg0) => (execute_EBREAK arg0) : M (Retired)
- | WFI (arg0) => (execute_WFI arg0) : M (Retired)
- | SFENCE_VMA ((rs1, rs2)) => (execute_SFENCE_VMA rs1 rs2) : M (Retired)
- | LOADRES ((aq, rl, rs1, width, rd)) => (execute_LOADRES aq rl rs1 width rd) : M (Retired)
- | STORECON ((aq, rl, rs2, rs1, width, rd)) =>
+ | STORE (imm, rs2, rs1, width, aq, rl) => (execute_STORE imm rs2 rs1 width aq rl) : M (Retired)
+ | ADDIW (imm, rs1, rd) => (execute_ADDIW imm rs1 rd) : M (Retired)
+ | SHIFTW (shamt, rs1, rd, op) => (execute_SHIFTW shamt rs1 rd op) : M (Retired)
+ | RTYPEW (rs2, rs1, rd, op) => (execute_RTYPEW rs2 rs1 rd op) : M (Retired)
+ | SHIFTIWOP (shamt, rs1, rd, op) => (execute_SHIFTIWOP shamt rs1 rd op) : M (Retired)
+ | FENCE (pred, succ) => (execute_FENCE pred succ) : M (Retired)
+ | FENCE_TSO (pred, succ) => (execute_FENCE_TSO pred succ) : M (Retired)
+ | FENCEI arg0 => returnm (execute_FENCEI arg0)
+ | ECALL arg0 => (execute_ECALL arg0) : M (Retired)
+ | MRET arg0 => (execute_MRET arg0) : M (Retired)
+ | SRET arg0 => (execute_SRET arg0) : M (Retired)
+ | EBREAK arg0 => (execute_EBREAK arg0) : M (Retired)
+ | WFI arg0 => (execute_WFI arg0) : M (Retired)
+ | SFENCE_VMA (rs1, rs2) => (execute_SFENCE_VMA rs1 rs2) : M (Retired)
+ | LOADRES (aq, rl, rs1, width, rd) => (execute_LOADRES aq rl rs1 width rd) : M (Retired)
+ | STORECON (aq, rl, rs2, rs1, width, rd) =>
(execute_STORECON aq rl rs2 rs1 width rd) : M (Retired)
- | AMO ((op, aq, rl, rs2, rs1, width, rd)) =>
+ | AMO (op, aq, rl, rs2, rs1, width, rd) =>
(execute_AMO op aq rl rs2 rs1 width rd) : M (Retired)
- | C_NOP (arg0) => returnm ((execute_C_NOP arg0) : Retired)
- | MUL ((rs2, rs1, rd, high, signed1, signed2)) =>
+ | C_NOP arg0 => returnm (execute_C_NOP arg0)
+ | MUL (rs2, rs1, rd, high, signed1, signed2) =>
(execute_MUL rs2 rs1 rd high signed1 signed2) : M (Retired)
- | DIV ((rs2, rs1, rd, s)) => (execute_DIV rs2 rs1 rd s) : M (Retired)
- | REM ((rs2, rs1, rd, s)) => (execute_REM rs2 rs1 rd s) : M (Retired)
- | MULW ((rs2, rs1, rd)) => (execute_MULW rs2 rs1 rd) : M (Retired)
- | DIVW ((rs2, rs1, rd, s)) => (execute_DIVW rs2 rs1 rd s) : M (Retired)
- | REMW ((rs2, rs1, rd, s)) => (execute_REMW rs2 rs1 rd s) : M (Retired)
- | CSR ((csr, rs1, rd, is_imm, op)) => (execute_CSR csr rs1 rd is_imm op) : M (Retired)
- | URET (arg0) => (execute_URET arg0) : M (Retired)
- | RISCV_JALR ((imm, rs1, rd)) => (execute_RISCV_JALR imm rs1 rd) : M (Retired)
- | ILLEGAL (s) => (execute_ILLEGAL s) : M (Retired)
- | C_ILLEGAL (s) => (execute_C_ILLEGAL s) : M (Retired)
+ | DIV (rs2, rs1, rd, s) => (execute_DIV rs2 rs1 rd s) : M (Retired)
+ | REM (rs2, rs1, rd, s) => (execute_REM rs2 rs1 rd s) : M (Retired)
+ | MULW (rs2, rs1, rd) => (execute_MULW rs2 rs1 rd) : M (Retired)
+ | DIVW (rs2, rs1, rd, s) => (execute_DIVW rs2 rs1 rd s) : M (Retired)
+ | REMW (rs2, rs1, rd, s) => (execute_REMW rs2 rs1 rd s) : M (Retired)
+ | CSR (csr, rs1, rd, is_imm, op) => (execute_CSR csr rs1 rd is_imm op) : M (Retired)
+ | URET arg0 => (execute_URET arg0) : M (Retired)
+ | LOAD_FP (imm, rs1, rd, width) => (execute_LOAD_FP imm rs1 rd width) : M (Retired)
+ | STORE_FP (imm, rs2, rs1, width) => (execute_STORE_FP imm rs2 rs1 width) : M (Retired)
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, op) =>
+ (execute_F_MADD_TYPE_S rs3 rs2 rs1 rm rd op) : M (Retired)
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, op) =>
+ (execute_F_BIN_RM_TYPE_S rs2 rs1 rm rd op) : M (Retired)
+ | F_UN_RM_TYPE_S (rs1, rm, rd, arg3) => (execute_F_UN_RM_TYPE_S rs1 rm rd arg3) : M (Retired)
+ | F_BIN_TYPE_S (rs2, rs1, rd, arg3) => (execute_F_BIN_TYPE_S rs2 rs1 rd arg3) : M (Retired)
+ | F_UN_TYPE_S (rs1, rd, arg2) => (execute_F_UN_TYPE_S rs1 rd arg2) : M (Retired)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, op) =>
+ (execute_F_MADD_TYPE_D rs3 rs2 rs1 rm rd op) : M (Retired)
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, op) =>
+ (execute_F_BIN_RM_TYPE_D rs2 rs1 rm rd op) : M (Retired)
+ | F_UN_RM_TYPE_D (rs1, rm, rd, arg3) => (execute_F_UN_RM_TYPE_D rs1 rm rd arg3) : M (Retired)
+ | F_BIN_TYPE_D (rs2, rs1, rd, arg3) => (execute_F_BIN_TYPE_D rs2 rs1 rd arg3) : M (Retired)
+ | F_UN_TYPE_D (rs1, rd, arg2) => (execute_F_UN_TYPE_D rs1 rd arg2) : M (Retired)
+ | RISCV_JALR (imm, rs1, rd) => (execute_RISCV_JALR imm rs1 rd) : M (Retired)
+ | ILLEGAL s => (execute_ILLEGAL s) : M (Retired)
+ | C_ILLEGAL s => (execute_C_ILLEGAL s) : M (Retired)
end)
- : M (Retired).
+ : M (Retired)
+).
+Defined.
-Definition execute (i : ast)
-: M (Retired) :=
-
- (_rec_execute i ((compressed_measure i) : Z) (Zwf_guarded _))
- : M (Retired).
-Definition assembly_forwards (arg_ : ast)
-: M (string) :=
-
+Definition execute (i : ast) : M (Retired) :=
+ (_rec_execute i ((compressed_measure i) : Z) (Zwf_guarded _)) : M (Retired).
+
+Definition assembly_forwards (arg_ : ast) : M (string) :=
(match arg_ with
- | UTYPE ((imm, rd, op)) =>
+ | UTYPE (imm, rd, op) =>
(reg_name_forwards rd) >>= fun w__0 : string =>
- returnm ((string_append (utype_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__0
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))
- : string)
- | RISCV_JAL ((imm, rd)) =>
+ returnm (string_append (utype_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__0
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
+ | RISCV_JAL (imm, rd) =>
(reg_name_forwards rd) >>= fun w__1 : string =>
- returnm ((string_append "jal"
- (string_append (spc_forwards tt)
- (string_append w__1
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))
- : string)
- | RISCV_JALR ((imm, rs1, rd)) =>
+ returnm (string_append "jal"
+ (string_append (spc_forwards tt)
+ (string_append w__1
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
+ | RISCV_JALR (imm, rs1, rd) =>
(reg_name_forwards rd) >>= fun w__2 : string =>
(reg_name_forwards rs1) >>= fun w__3 : string =>
- returnm ((string_append "jalr"
- (string_append (spc_forwards tt)
- (string_append w__2
- (string_append (sep_forwards tt)
- (string_append w__3
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))))
- : string)
- | BTYPE ((imm, rs2, rs1, op)) =>
+ returnm (string_append "jalr"
+ (string_append (spc_forwards tt)
+ (string_append w__2
+ (string_append (sep_forwards tt)
+ (string_append w__3
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))))
+ | BTYPE (imm, rs2, rs1, op) =>
(reg_name_forwards rs1) >>= fun w__4 : string =>
(reg_name_forwards rs2) >>= fun w__5 : string =>
- returnm ((string_append (btype_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__4
- (string_append (sep_forwards tt)
- (string_append w__5
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))))
- : string)
- | ITYPE ((imm, rs1, rd, op)) =>
+ returnm (string_append (btype_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__4
+ (string_append (sep_forwards tt)
+ (string_append w__5
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))))
+ | ITYPE (imm, rs1, rd, op) =>
(reg_name_forwards rd) >>= fun w__6 : string =>
(reg_name_forwards rs1) >>= fun w__7 : string =>
- returnm ((string_append (itype_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__6
- (string_append (sep_forwards tt)
- (string_append w__7
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))))
- : string)
- | SHIFTIOP ((shamt, rs1, rd, op)) =>
+ returnm (string_append (itype_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__6
+ (string_append (sep_forwards tt)
+ (string_append w__7
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))))
+ | SHIFTIOP (shamt, rs1, rd, op) =>
(reg_name_forwards rd) >>= fun w__8 : string =>
(reg_name_forwards rs1) >>= fun w__9 : string =>
- returnm ((string_append (shiftiop_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__8
- (string_append (sep_forwards tt)
- (string_append w__9
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits shamt) "")))))))
- : string)
- | RTYPE ((rs2, rs1, rd, op)) =>
+ returnm (string_append (shiftiop_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__8
+ (string_append (sep_forwards tt)
+ (string_append w__9
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits shamt) "")))))))
+ | RTYPE (rs2, rs1, rd, op) =>
(reg_name_forwards rd) >>= fun w__10 : string =>
(reg_name_forwards rs1) >>= fun w__11 : string =>
(reg_name_forwards rs2) >>= fun w__12 : string =>
- returnm ((string_append (rtype_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__10
- (string_append (sep_forwards tt)
- (string_append w__11
- (string_append (sep_forwards tt) (string_append w__12 "")))))))
- : string)
- | LOAD ((imm, rs1, rd, is_unsigned, size, aq, rl)) =>
+ returnm (string_append (rtype_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__10
+ (string_append (sep_forwards tt)
+ (string_append w__11
+ (string_append (sep_forwards tt) (string_append w__12 "")))))))
+ | LOAD (imm, rs1, rd, is_unsigned, size, aq, rl) =>
(reg_name_forwards rd) >>= fun w__13 : string =>
(reg_name_forwards rs1) >>= fun w__14 : string =>
- returnm ((string_append "l"
- (string_append (size_mnemonic_forwards size)
- (string_append (maybe_u_forwards is_unsigned)
- (string_append (maybe_aq_forwards aq)
- (string_append (maybe_rl_forwards rl)
- (string_append (spc_forwards tt)
- (string_append w__13
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm)
- (string_append (opt_spc_forwards tt)
- (string_append "("
- (string_append (opt_spc_forwards tt)
- (string_append w__14
- (string_append (opt_spc_forwards tt)
- (string_append ")" "")))))))))))))))
- : string)
- | STORE ((imm, rs2, rs1, size, aq, rl)) =>
+ returnm (string_append "l"
+ (string_append (size_mnemonic_forwards size)
+ (string_append (maybe_u_forwards is_unsigned)
+ (string_append (maybe_aq_forwards aq)
+ (string_append (maybe_rl_forwards rl)
+ (string_append (spc_forwards tt)
+ (string_append w__13
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm)
+ (string_append (opt_spc_forwards tt)
+ (string_append "("
+ (string_append (opt_spc_forwards tt)
+ (string_append w__14
+ (string_append (opt_spc_forwards tt)
+ (string_append ")" "")))))))))))))))
+ | STORE (imm, rs2, rs1, size, aq, rl) =>
(reg_name_forwards rs2) >>= fun w__15 : string =>
(reg_name_forwards rs1) >>= fun w__16 : string =>
- returnm ((string_append "s"
- (string_append (size_mnemonic_forwards size)
- (string_append (maybe_aq_forwards aq)
- (string_append (maybe_rl_forwards rl)
- (string_append (spc_forwards tt)
- (string_append w__15
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm)
- (string_append (opt_spc_forwards tt)
- (string_append "("
- (string_append (opt_spc_forwards tt)
- (string_append w__16
- (string_append (opt_spc_forwards tt)
- (string_append ")" ""))))))))))))))
- : string)
- | ADDIW ((imm, rs1, rd)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ returnm (string_append "s"
+ (string_append (size_mnemonic_forwards size)
+ (string_append (maybe_aq_forwards aq)
+ (string_append (maybe_rl_forwards rl)
+ (string_append (spc_forwards tt)
+ (string_append w__15
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm)
+ (string_append (opt_spc_forwards tt)
+ (string_append "("
+ (string_append (opt_spc_forwards tt)
+ (string_append w__16
+ (string_append (opt_spc_forwards tt)
+ (string_append ")" ""))))))))))))))
+ | ADDIW (imm, rs1, rd) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rd) >>= fun w__17 : string =>
(reg_name_forwards rs1) >>= fun w__18 : string =>
- returnm ((string_append "addiw"
- (string_append (spc_forwards tt)
- (string_append w__17
- (string_append (sep_forwards tt)
- (string_append w__18
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))))
- : string)
+ returnm (string_append "addiw"
+ (string_append (spc_forwards tt)
+ (string_append w__17
+ (string_append (sep_forwards tt)
+ (string_append w__18
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | SHIFTW ((shamt, rs1, rd, op)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ | SHIFTW (shamt, rs1, rd, op) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rd) >>= fun w__21 : string =>
(reg_name_forwards rs1) >>= fun w__22 : string =>
- returnm ((string_append (shiftw_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__21
- (string_append (sep_forwards tt)
- (string_append w__22
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits shamt) "")))))))
- : string)
+ returnm (string_append (shiftw_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__21
+ (string_append (sep_forwards tt)
+ (string_append w__22
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits shamt) "")))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | RTYPEW ((rs2, rs1, rd, op)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ | RTYPEW (rs2, rs1, rd, op) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rd) >>= fun w__25 : string =>
(reg_name_forwards rs1) >>= fun w__26 : string =>
(reg_name_forwards rs2) >>= fun w__27 : string =>
- returnm ((string_append (rtypew_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__25
- (string_append (sep_forwards tt)
- (string_append w__26
- (string_append (sep_forwards tt) (string_append w__27 "")))))))
- : string)
+ returnm (string_append (rtypew_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__25
+ (string_append (sep_forwards tt)
+ (string_append w__26
+ (string_append (sep_forwards tt) (string_append w__27 "")))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | SHIFTIWOP ((shamt, rs1, rd, op)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ | SHIFTIWOP (shamt, rs1, rd, op) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rd) >>= fun w__30 : string =>
(reg_name_forwards rs1) >>= fun w__31 : string =>
- returnm ((string_append (shiftiwop_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__30
- (string_append (sep_forwards tt)
- (string_append w__31
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits shamt) "")))))))
- : string)
+ returnm (string_append (shiftiwop_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__30
+ (string_append (sep_forwards tt)
+ (string_append w__31
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits shamt) "")))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | FENCE ((pred, succ)) =>
+ | FENCE (pred, succ) =>
(fence_bits_forwards pred) >>= fun w__34 : string =>
(fence_bits_forwards succ) >>= fun w__35 : string =>
- returnm ((string_append "fence"
- (string_append (spc_forwards tt)
- (string_append w__34
- (string_append (sep_forwards tt) (string_append w__35 "")))))
- : string)
- | FENCE_TSO ((pred, succ)) =>
+ returnm (string_append "fence"
+ (string_append (spc_forwards tt)
+ (string_append w__34 (string_append (sep_forwards tt) (string_append w__35 "")))))
+ | FENCE_TSO (pred, succ) =>
(fence_bits_forwards pred) >>= fun w__36 : string =>
(fence_bits_forwards succ) >>= fun w__37 : string =>
- returnm ((string_append "fence.tso"
- (string_append (spc_forwards tt)
- (string_append w__36
- (string_append (sep_forwards tt) (string_append w__37 "")))))
- : string)
- | FENCEI (tt) => returnm ("fence.i" : string)
- | ECALL (tt) => returnm ("ecall" : string)
- | MRET (tt) => returnm ("mret" : string)
- | SRET (tt) => returnm ("sret" : string)
- | EBREAK (tt) => returnm ("ebreak" : string)
- | WFI (tt) => returnm ("wfi" : string)
- | SFENCE_VMA ((rs1, rs2)) =>
+ returnm (string_append "fence.tso"
+ (string_append (spc_forwards tt)
+ (string_append w__36 (string_append (sep_forwards tt) (string_append w__37 "")))))
+ | FENCEI tt => returnm "fence.i"
+ | ECALL tt => returnm "ecall"
+ | MRET tt => returnm "mret"
+ | SRET tt => returnm "sret"
+ | EBREAK tt => returnm "ebreak"
+ | WFI tt => returnm "wfi"
+ | SFENCE_VMA (rs1, rs2) =>
(reg_name_forwards rs1) >>= fun w__38 : string =>
(reg_name_forwards rs2) >>= fun w__39 : string =>
- returnm ((string_append "sfence.vma"
- (string_append (spc_forwards tt)
- (string_append w__38
- (string_append (sep_forwards tt) (string_append w__39 "")))))
- : string)
- | LOADRES ((aq, rl, rs1, size, rd)) =>
+ returnm (string_append "sfence.vma"
+ (string_append (spc_forwards tt)
+ (string_append w__38 (string_append (sep_forwards tt) (string_append w__39 "")))))
+ | LOADRES (aq, rl, rs1, size, rd) =>
(reg_name_forwards rd) >>= fun w__40 : string =>
(reg_name_forwards rs1) >>= fun w__41 : string =>
- returnm ((string_append "lr."
- (string_append (size_mnemonic_forwards size)
- (string_append (maybe_aq_forwards aq)
- (string_append (maybe_rl_forwards rl)
- (string_append (spc_forwards tt)
- (string_append w__40
- (string_append (sep_forwards tt) (string_append w__41 ""))))))))
- : string)
- | STORECON ((aq, rl, rs2, rs1, size, rd)) =>
+ returnm (string_append "lr."
+ (string_append (size_mnemonic_forwards size)
+ (string_append (maybe_aq_forwards aq)
+ (string_append (maybe_rl_forwards rl)
+ (string_append (spc_forwards tt)
+ (string_append w__40
+ (string_append (sep_forwards tt) (string_append w__41 ""))))))))
+ | STORECON (aq, rl, rs2, rs1, size, rd) =>
(reg_name_forwards rd) >>= fun w__42 : string =>
(reg_name_forwards rs1) >>= fun w__43 : string =>
(reg_name_forwards rs2) >>= fun w__44 : string =>
- returnm ((string_append "sc."
- (string_append (size_mnemonic_forwards size)
- (string_append (maybe_aq_forwards aq)
- (string_append (maybe_rl_forwards rl)
- (string_append (spc_forwards tt)
- (string_append w__42
- (string_append (sep_forwards tt)
- (string_append w__43
- (string_append (sep_forwards tt) (string_append w__44 ""))))))))))
- : string)
- | AMO ((op, aq, rl, rs2, rs1, width, rd)) =>
+ returnm (string_append "sc."
+ (string_append (size_mnemonic_forwards size)
+ (string_append (maybe_aq_forwards aq)
+ (string_append (maybe_rl_forwards rl)
+ (string_append (spc_forwards tt)
+ (string_append w__42
+ (string_append (sep_forwards tt)
+ (string_append w__43
+ (string_append (sep_forwards tt) (string_append w__44 ""))))))))))
+ | AMO (op, aq, rl, rs2, rs1, width, rd) =>
(reg_name_forwards rd) >>= fun w__45 : string =>
- (reg_name_forwards rs1) >>= fun w__46 : string =>
- (reg_name_forwards rs2) >>= fun w__47 : string =>
- returnm ((string_append (amo_mnemonic_forwards op)
- (string_append "."
- (string_append (size_mnemonic_forwards width)
- (string_append (maybe_aq_forwards aq)
- (string_append (maybe_rl_forwards rl)
- (string_append (spc_forwards tt)
- (string_append w__45
- (string_append (sep_forwards tt)
- (string_append w__46
- (string_append (sep_forwards tt) (string_append w__47 "")))))))))))
- : string)
- | C_NOP (tt) => returnm ("c.nop" : string)
- | C_ADDI4SPN ((rdc, nzimm)) =>
- (if ((neq_vec nzimm (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8))) then
+ (reg_name_forwards rs2) >>= fun w__46 : string =>
+ (reg_name_forwards rs1) >>= fun w__47 : string =>
+ returnm (string_append (amo_mnemonic_forwards op)
+ (string_append "."
+ (string_append (size_mnemonic_forwards width)
+ (string_append (maybe_aq_forwards aq)
+ (string_append (maybe_rl_forwards rl)
+ (string_append (spc_forwards tt)
+ (string_append w__45
+ (string_append (sep_forwards tt)
+ (string_append w__46
+ (string_append (sep_forwards tt)
+ (string_append "("
+ (string_append w__47 (string_append ")" "")))))))))))))
+ | C_NOP tt => returnm "c.nop"
+ | C_ADDI4SPN (rdc, nzimm) =>
+ (if neq_vec nzimm (Ox"00" : mword 8) then
(creg_name_forwards rdc) >>= fun w__48 : string =>
- returnm ((string_append "c.addi4spn"
- (string_append (spc_forwards tt)
- (string_append w__48
- (string_append (sep_forwards tt)
- (string_append
- (decimal_string_of_bits
- (concat_vec (nzimm : mword 8) (vec_of_bits [B0;B0] : mword 2)))
- "")))))
- : string)
+ returnm (string_append "c.addi4spn"
+ (string_append (spc_forwards tt)
+ (string_append w__48
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (nzimm : mword 8) ('b"00" : mword 2))) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_LW ((uimm, rsc, rdc)) =>
+ | C_LW (uimm, rsc, rdc) =>
(creg_name_forwards rdc) >>= fun w__51 : string =>
(creg_name_forwards rsc) >>= fun w__52 : string =>
- returnm ((string_append "c.lw"
- (string_append (spc_forwards tt)
- (string_append w__51
- (string_append (sep_forwards tt)
- (string_append w__52
- (string_append (sep_forwards tt)
- (string_append
- (decimal_string_of_bits
- (concat_vec (uimm : mword 5)
- (vec_of_bits [B0;B0] : mword 2))) "")))))))
- : string)
- | C_LD ((uimm, rsc, rdc)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ returnm (string_append "c.lw"
+ (string_append (spc_forwards tt)
+ (string_append w__51
+ (string_append (sep_forwards tt)
+ (string_append w__52
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (uimm : mword 5) ('b"00" : mword 2))) "")))))))
+ | C_LD (uimm, rsc, rdc) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(creg_name_forwards rdc) >>= fun w__53 : string =>
(creg_name_forwards rsc) >>= fun w__54 : string =>
- returnm ((string_append "c.ld"
- (string_append (spc_forwards tt)
- (string_append w__53
- (string_append (sep_forwards tt)
- (string_append w__54
- (string_append (sep_forwards tt)
- (string_append
- (decimal_string_of_bits
- (concat_vec (uimm : mword 5)
- (vec_of_bits [B0;B0;B0] : mword 3))) "")))))))
- : string)
+ returnm (string_append "c.ld"
+ (string_append (spc_forwards tt)
+ (string_append w__53
+ (string_append (sep_forwards tt)
+ (string_append w__54
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (uimm : mword 5) ('b"000" : mword 3))) "")))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_SW ((uimm, rsc1, rsc2)) =>
+ | C_SW (uimm, rsc1, rsc2) =>
(creg_name_forwards rsc1) >>= fun w__57 : string =>
(creg_name_forwards rsc2) >>= fun w__58 : string =>
- returnm ((string_append "c.sw"
- (string_append (spc_forwards tt)
- (string_append w__57
- (string_append (sep_forwards tt)
- (string_append w__58
- (string_append (sep_forwards tt)
- (string_append
- (decimal_string_of_bits
- (concat_vec (uimm : mword 5)
- (vec_of_bits [B0;B0] : mword 2))) "")))))))
- : string)
- | C_SD ((uimm, rsc1, rsc2)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ returnm (string_append "c.sw"
+ (string_append (spc_forwards tt)
+ (string_append w__57
+ (string_append (sep_forwards tt)
+ (string_append w__58
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (uimm : mword 5) ('b"00" : mword 2))) "")))))))
+ | C_SD (uimm, rsc1, rsc2) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(creg_name_forwards rsc1) >>= fun w__59 : string =>
(creg_name_forwards rsc2) >>= fun w__60 : string =>
- returnm ((string_append "c.sd"
- (string_append (spc_forwards tt)
- (string_append w__59
- (string_append (sep_forwards tt)
- (string_append w__60
- (string_append (sep_forwards tt)
- (string_append
- (decimal_string_of_bits
- (concat_vec (uimm : mword 5)
- (vec_of_bits [B0;B0;B0] : mword 3))) "")))))))
- : string)
+ returnm (string_append "c.sd"
+ (string_append (spc_forwards tt)
+ (string_append w__59
+ (string_append (sep_forwards tt)
+ (string_append w__60
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (uimm : mword 5) ('b"000" : mword 3))) "")))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_ADDI ((nzi, rsd)) =>
- (if ((andb (neq_vec nzi (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- : bool))) then
+ | C_ADDI (nzi, rsd) =>
+ (if andb (neq_vec nzi ('b"000000" : mword 6)) (neq_vec rsd zreg) then
(reg_name_forwards rsd) >>= fun w__63 : string =>
- returnm ((string_append "c.addi"
- (string_append (spc_forwards tt)
- (string_append w__63
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits nzi) "")))))
- : string)
+ returnm (string_append "c.addi"
+ (string_append (spc_forwards tt)
+ (string_append w__63
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits nzi) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_JAL (imm) =>
- (if sumbool_of_bool ((Z.eqb 64 32)) then
- returnm ((string_append "c.jal"
- (string_append (spc_forwards tt)
- (string_append
- (decimal_string_of_bits
- (concat_vec (imm : mword 11) (vec_of_bits [B0] : mword 1))) "")))
- : string)
+ | C_JAL imm =>
+ (if sumbool_of_bool (Z.eqb 64 32) then
+ returnm (string_append "c.jal"
+ (string_append (spc_forwards tt)
+ (string_append
+ (decimal_string_of_bits (concat_vec (imm : mword 11) ('b"0" : mword 1)))
+ "")))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_ADDIW ((imm, rsd)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ | C_ADDIW (imm, rsd) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rsd) >>= fun w__68 : string =>
- returnm ((string_append "c.addiw"
- (string_append (spc_forwards tt)
- (string_append w__68
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))
- : string)
+ returnm (string_append "c.addiw"
+ (string_append (spc_forwards tt)
+ (string_append w__68
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_LI ((imm, rd)) =>
- (if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg)))))
- then
+ | C_LI (imm, rd) =>
+ (if neq_vec rd zreg then
(reg_name_forwards rd) >>= fun w__71 : string =>
- returnm ((string_append "c.li"
- (string_append (spc_forwards tt)
- (string_append w__71
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))
- : string)
+ returnm (string_append "c.li"
+ (string_append (spc_forwards tt)
+ (string_append w__71
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_ADDI16SP (imm) =>
- (if ((neq_vec imm (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then
- returnm ((string_append "c.addi16sp"
- (string_append (spc_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))
- : string)
+ | C_ADDI16SP imm =>
+ (if neq_vec imm ('b"000000" : mword 6) then
+ returnm (string_append "c.addi16sp"
+ (string_append (spc_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_LUI ((imm, rd)) =>
- (if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno sp))))
- (neq_vec imm (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)))
- : bool))) then
+ | C_LUI (imm, rd) =>
+ (if andb (neq_vec rd zreg) (andb (neq_vec rd sp) (neq_vec imm ('b"000000" : mword 6))) then
(reg_name_forwards rd) >>= fun w__76 : string =>
- returnm ((string_append "c.lui"
- (string_append (spc_forwards tt)
- (string_append w__76
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))
- : string)
+ returnm (string_append "c.lui"
+ (string_append (spc_forwards tt)
+ (string_append w__76
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_SRLI ((shamt, rsd)) =>
- (if ((neq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then
+ | C_SRLI (shamt, rsd) =>
+ (if neq_vec shamt ('b"000000" : mword 6) then
(creg_name_forwards rsd) >>= fun w__79 : string =>
- returnm ((string_append "c.srli"
- (string_append (spc_forwards tt)
- (string_append w__79
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits shamt) "")))))
- : string)
+ returnm (string_append "c.srli"
+ (string_append (spc_forwards tt)
+ (string_append w__79
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits shamt) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_SRAI ((shamt, rsd)) =>
- (if ((neq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then
+ | C_SRAI (shamt, rsd) =>
+ (if neq_vec shamt ('b"000000" : mword 6) then
(creg_name_forwards rsd) >>= fun w__82 : string =>
- returnm ((string_append "c.srai"
- (string_append (spc_forwards tt)
- (string_append w__82
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits shamt) "")))))
- : string)
+ returnm (string_append "c.srai"
+ (string_append (spc_forwards tt)
+ (string_append w__82
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits shamt) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_ANDI ((imm, rsd)) =>
+ | C_ANDI (imm, rsd) =>
(creg_name_forwards rsd) >>= fun w__85 : string =>
- returnm ((string_append "c.andi"
- (string_append (spc_forwards tt)
- (string_append w__85
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))
- : string)
- | C_SUB ((rsd, rs2)) =>
+ returnm (string_append "c.andi"
+ (string_append (spc_forwards tt)
+ (string_append w__85
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
+ | C_SUB (rsd, rs2) =>
(creg_name_forwards rsd) >>= fun w__86 : string =>
(creg_name_forwards rs2) >>= fun w__87 : string =>
- returnm ((string_append "c.sub"
- (string_append (spc_forwards tt)
- (string_append w__86
- (string_append (sep_forwards tt) (string_append w__87 "")))))
- : string)
- | C_XOR ((rsd, rs2)) =>
+ returnm (string_append "c.sub"
+ (string_append (spc_forwards tt)
+ (string_append w__86 (string_append (sep_forwards tt) (string_append w__87 "")))))
+ | C_XOR (rsd, rs2) =>
(creg_name_forwards rsd) >>= fun w__88 : string =>
(creg_name_forwards rs2) >>= fun w__89 : string =>
- returnm ((string_append "c.xor"
- (string_append (spc_forwards tt)
- (string_append w__88
- (string_append (sep_forwards tt) (string_append w__89 "")))))
- : string)
- | C_OR ((rsd, rs2)) =>
+ returnm (string_append "c.xor"
+ (string_append (spc_forwards tt)
+ (string_append w__88 (string_append (sep_forwards tt) (string_append w__89 "")))))
+ | C_OR (rsd, rs2) =>
(creg_name_forwards rsd) >>= fun w__90 : string =>
(creg_name_forwards rs2) >>= fun w__91 : string =>
- returnm ((string_append "c.or"
- (string_append (spc_forwards tt)
- (string_append w__90
- (string_append (sep_forwards tt) (string_append w__91 "")))))
- : string)
- | C_AND ((rsd, rs2)) =>
+ returnm (string_append "c.or"
+ (string_append (spc_forwards tt)
+ (string_append w__90 (string_append (sep_forwards tt) (string_append w__91 "")))))
+ | C_AND (rsd, rs2) =>
(creg_name_forwards rsd) >>= fun w__92 : string =>
(creg_name_forwards rs2) >>= fun w__93 : string =>
- returnm ((string_append "c.and"
- (string_append (spc_forwards tt)
- (string_append w__92
- (string_append (sep_forwards tt) (string_append w__93 "")))))
- : string)
- | C_SUBW ((rsd, rs2)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ returnm (string_append "c.and"
+ (string_append (spc_forwards tt)
+ (string_append w__92 (string_append (sep_forwards tt) (string_append w__93 "")))))
+ | C_SUBW (rsd, rs2) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(creg_name_forwards rsd) >>= fun w__94 : string =>
(creg_name_forwards rs2) >>= fun w__95 : string =>
- returnm ((string_append "c.subw"
- (string_append (spc_forwards tt)
- (string_append w__94
- (string_append (sep_forwards tt) (string_append w__95 "")))))
- : string)
+ returnm (string_append "c.subw"
+ (string_append (spc_forwards tt)
+ (string_append w__94
+ (string_append (sep_forwards tt) (string_append w__95 "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_ADDW ((rsd, rs2)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ | C_ADDW (rsd, rs2) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(creg_name_forwards rsd) >>= fun w__98 : string =>
(creg_name_forwards rs2) >>= fun w__99 : string =>
- returnm ((string_append "c.addw"
- (string_append (spc_forwards tt)
- (string_append w__98
- (string_append (sep_forwards tt) (string_append w__99 "")))))
- : string)
+ returnm (string_append "c.addw"
+ (string_append (spc_forwards tt)
+ (string_append w__98
+ (string_append (sep_forwards tt) (string_append w__99 "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_J (imm) =>
- returnm ((string_append "c.j"
- (string_append (spc_forwards tt) (string_append (decimal_string_of_bits imm) "")))
- : string)
- | C_BEQZ ((imm, rs)) =>
+ | C_J imm =>
+ returnm (string_append "c.j"
+ (string_append (spc_forwards tt) (string_append (decimal_string_of_bits imm) "")))
+ | C_BEQZ (imm, rs) =>
(creg_name_forwards rs) >>= fun w__102 : string =>
- returnm ((string_append "c.beqz"
- (string_append (spc_forwards tt)
- (string_append w__102
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))
- : string)
- | C_BNEZ ((imm, rs)) =>
+ returnm (string_append "c.beqz"
+ (string_append (spc_forwards tt)
+ (string_append w__102
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
+ | C_BNEZ (imm, rs) =>
(creg_name_forwards rs) >>= fun w__103 : string =>
- returnm ((string_append "c.bnez"
- (string_append (spc_forwards tt)
- (string_append w__103
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits imm) "")))))
- : string)
- | C_SLLI ((shamt, rsd)) =>
- (if ((andb (neq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- : bool))) then
+ returnm (string_append "c.bnez"
+ (string_append (spc_forwards tt)
+ (string_append w__103
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
+ | C_SLLI (shamt, rsd) =>
+ (if andb (neq_vec shamt ('b"000000" : mword 6)) (neq_vec rsd zreg) then
(reg_name_forwards rsd) >>= fun w__104 : string =>
- returnm ((string_append "c.slli"
- (string_append (spc_forwards tt)
- (string_append w__104
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits shamt) "")))))
- : string)
+ returnm (string_append "c.slli"
+ (string_append (spc_forwards tt)
+ (string_append w__104
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits shamt) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_LWSP ((uimm, rd)) =>
- (if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg)))))
- then
+ | C_LWSP (uimm, rd) =>
+ (if neq_vec rd zreg then
(reg_name_forwards rd) >>= fun w__107 : string =>
- returnm ((string_append "c.lwsp"
- (string_append (spc_forwards tt)
- (string_append w__107
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits uimm) "")))))
- : string)
+ returnm (string_append "c.lwsp"
+ (string_append (spc_forwards tt)
+ (string_append w__107
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits uimm) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_LDSP ((uimm, rd)) =>
- (if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg)))) (Z.eqb 64 64))) then
+ | C_LDSP (uimm, rd) =>
+ (if sumbool_of_bool (andb (neq_vec rd zreg) (Z.eqb 64 64)) then
(reg_name_forwards rd) >>= fun w__110 : string =>
- returnm ((string_append "c.ldsp"
- (string_append (spc_forwards tt)
- (string_append w__110
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits uimm) "")))))
- : string)
+ returnm (string_append "c.ldsp"
+ (string_append (spc_forwards tt)
+ (string_append w__110
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits uimm) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_SWSP ((uimm, rd)) =>
+ | C_SWSP (uimm, rd) =>
(reg_name_forwards rd) >>= fun w__113 : string =>
- returnm ((string_append "c.swsp"
- (string_append (spc_forwards tt)
- (string_append w__113
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits uimm) "")))))
- : string)
- | C_SDSP ((uimm, rs2)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ returnm (string_append "c.swsp"
+ (string_append (spc_forwards tt)
+ (string_append w__113
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits uimm) "")))))
+ | C_SDSP (uimm, rs2) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rs2) >>= fun w__114 : string =>
- returnm ((string_append "c.sdsp"
- (string_append (spc_forwards tt)
- (string_append w__114
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits uimm) "")))))
- : string)
+ returnm (string_append "c.sdsp"
+ (string_append (spc_forwards tt)
+ (string_append w__114
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits uimm) "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_JR (rs1) =>
- (if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rs1)) (projT1 (regidx_to_regno zreg)))))
- then
+ | C_JR rs1 =>
+ (if neq_vec rs1 zreg then
(reg_name_forwards rs1) >>= fun w__117 : string =>
- returnm ((string_append "c.jr" (string_append (spc_forwards tt) (string_append w__117 "")))
- : string)
+ returnm (string_append "c.jr" (string_append (spc_forwards tt) (string_append w__117 "")))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_JALR (rs1) =>
- (if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rs1)) (projT1 (regidx_to_regno zreg)))))
- then
+ | C_JALR rs1 =>
+ (if neq_vec rs1 zreg then
(reg_name_forwards rs1) >>= fun w__120 : string =>
- returnm ((string_append "c.jalr"
- (string_append (spc_forwards tt) (string_append w__120 "")))
- : string)
+ returnm (string_append "c.jalr"
+ (string_append (spc_forwards tt) (string_append w__120 "")))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_MV ((rd, rs2)) =>
- (if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))) then
+ | C_MV (rd, rs2) =>
+ (if andb (neq_vec rd zreg) (neq_vec rs2 zreg) then
(reg_name_forwards rd) >>= fun w__123 : string =>
(reg_name_forwards rs2) >>= fun w__124 : string =>
- returnm ((string_append "c.mv"
- (string_append (spc_forwards tt)
- (string_append w__123
- (string_append (sep_forwards tt) (string_append w__124 "")))))
- : string)
+ returnm (string_append "c.mv"
+ (string_append (spc_forwards tt)
+ (string_append w__123
+ (string_append (sep_forwards tt) (string_append w__124 "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | C_EBREAK (tt) => returnm ("c.ebreak" : string)
- | C_ADD ((rsd, rs2)) =>
- (if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
- (projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))) then
+ | C_EBREAK tt => returnm "c.ebreak"
+ | C_ADD (rsd, rs2) =>
+ (if andb (neq_vec rsd zreg) (neq_vec rs2 zreg) then
(reg_name_forwards rsd) >>= fun w__127 : string =>
(reg_name_forwards rs2) >>= fun w__128 : string =>
- returnm ((string_append "c.add"
- (string_append (spc_forwards tt)
- (string_append w__127
- (string_append (sep_forwards tt) (string_append w__128 "")))))
- : string)
+ returnm (string_append "c.add"
+ (string_append (spc_forwards tt)
+ (string_append w__127
+ (string_append (sep_forwards tt) (string_append w__128 "")))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | MUL ((rs2, rs1, rd, high, signed1, signed2)) =>
+ | MUL (rs2, rs1, rd, high, signed1, signed2) =>
(mul_mnemonic_forwards (high, signed1, signed2)) >>= fun w__131 : string =>
(reg_name_forwards rd) >>= fun w__132 : string =>
(reg_name_forwards rs1) >>= fun w__133 : string =>
(reg_name_forwards rs2) >>= fun w__134 : string =>
- returnm ((string_append w__131
- (string_append (spc_forwards tt)
- (string_append w__132
- (string_append (sep_forwards tt)
- (string_append w__133
- (string_append (sep_forwards tt) (string_append w__134 "")))))))
- : string)
- | DIV ((rs2, rs1, rd, s)) =>
+ returnm (string_append w__131
+ (string_append (spc_forwards tt)
+ (string_append w__132
+ (string_append (sep_forwards tt)
+ (string_append w__133
+ (string_append (sep_forwards tt) (string_append w__134 "")))))))
+ | DIV (rs2, rs1, rd, s) =>
(reg_name_forwards rd) >>= fun w__135 : string =>
(reg_name_forwards rs1) >>= fun w__136 : string =>
(reg_name_forwards rs2) >>= fun w__137 : string =>
- returnm ((string_append "div"
- (string_append (maybe_not_u_forwards s)
- (string_append (spc_forwards tt)
- (string_append w__135
- (string_append (sep_forwards tt)
- (string_append w__136
- (string_append (sep_forwards tt) (string_append w__137 ""))))))))
- : string)
- | REM ((rs2, rs1, rd, s)) =>
+ returnm (string_append "div"
+ (string_append (maybe_not_u_forwards s)
+ (string_append (spc_forwards tt)
+ (string_append w__135
+ (string_append (sep_forwards tt)
+ (string_append w__136
+ (string_append (sep_forwards tt) (string_append w__137 ""))))))))
+ | REM (rs2, rs1, rd, s) =>
(reg_name_forwards rd) >>= fun w__138 : string =>
(reg_name_forwards rs1) >>= fun w__139 : string =>
(reg_name_forwards rs2) >>= fun w__140 : string =>
- returnm ((string_append "rem"
- (string_append (maybe_not_u_forwards s)
- (string_append (spc_forwards tt)
- (string_append w__138
- (string_append (sep_forwards tt)
- (string_append w__139
- (string_append (sep_forwards tt) (string_append w__140 ""))))))))
- : string)
- | MULW ((rs2, rs1, rd)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ returnm (string_append "rem"
+ (string_append (maybe_not_u_forwards s)
+ (string_append (spc_forwards tt)
+ (string_append w__138
+ (string_append (sep_forwards tt)
+ (string_append w__139
+ (string_append (sep_forwards tt) (string_append w__140 ""))))))))
+ | MULW (rs2, rs1, rd) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rd) >>= fun w__141 : string =>
(reg_name_forwards rs1) >>= fun w__142 : string =>
(reg_name_forwards rs2) >>= fun w__143 : string =>
- returnm ((string_append "mulw"
- (string_append (spc_forwards tt)
- (string_append w__141
- (string_append (sep_forwards tt)
- (string_append w__142
- (string_append (sep_forwards tt) (string_append w__143 "")))))))
- : string)
+ returnm (string_append "mulw"
+ (string_append (spc_forwards tt)
+ (string_append w__141
+ (string_append (sep_forwards tt)
+ (string_append w__142
+ (string_append (sep_forwards tt) (string_append w__143 "")))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | DIVW ((rs2, rs1, rd, s)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ | DIVW (rs2, rs1, rd, s) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rd) >>= fun w__146 : string =>
(reg_name_forwards rs1) >>= fun w__147 : string =>
(reg_name_forwards rs2) >>= fun w__148 : string =>
- returnm ((string_append "div"
- (string_append (maybe_not_u_forwards s)
- (string_append "w"
- (string_append (spc_forwards tt)
- (string_append w__146
- (string_append (sep_forwards tt)
- (string_append w__147
- (string_append (sep_forwards tt) (string_append w__148 "")))))))))
- : string)
+ returnm (string_append "div"
+ (string_append (maybe_not_u_forwards s)
+ (string_append "w"
+ (string_append (spc_forwards tt)
+ (string_append w__146
+ (string_append (sep_forwards tt)
+ (string_append w__147
+ (string_append (sep_forwards tt) (string_append w__148 "")))))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | REMW ((rs2, rs1, rd, s)) =>
- (if sumbool_of_bool ((Z.eqb 64 64)) then
+ | REMW (rs2, rs1, rd, s) =>
+ (if sumbool_of_bool (Z.eqb 64 64) then
(reg_name_forwards rd) >>= fun w__151 : string =>
(reg_name_forwards rs1) >>= fun w__152 : string =>
(reg_name_forwards rs2) >>= fun w__153 : string =>
- returnm ((string_append "rem"
- (string_append (maybe_not_u_forwards s)
- (string_append "w"
- (string_append (spc_forwards tt)
- (string_append w__151
- (string_append (sep_forwards tt)
- (string_append w__152
- (string_append (sep_forwards tt) (string_append w__153 "")))))))))
- : string)
+ returnm (string_append "rem"
+ (string_append (maybe_not_u_forwards s)
+ (string_append "w"
+ (string_append (spc_forwards tt)
+ (string_append w__151
+ (string_append (sep_forwards tt)
+ (string_append w__152
+ (string_append (sep_forwards tt) (string_append w__153 "")))))))))
else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
: M (string)
- | CSR ((csr, rs1, rd, true, op)) =>
+ | CSR (csr, rs1, rd, true, op) =>
(reg_name_forwards rd) >>= fun w__156 : string =>
- (csr_name_map_forwards csr) >>= fun w__157 : string =>
- returnm ((string_append (csr_mnemonic_forwards op)
- (string_append "i"
- (string_append (spc_forwards tt)
- (string_append w__156
- (string_append (sep_forwards tt)
- (string_append (decimal_string_of_bits rs1)
- (string_append (sep_forwards tt) (string_append w__157 ""))))))))
- : string)
- | CSR ((csr, rs1, rd, false, op)) =>
- (reg_name_forwards rd) >>= fun w__158 : string =>
- (reg_name_forwards rs1) >>= fun w__159 : string =>
- (csr_name_map_forwards csr) >>= fun w__160 : string =>
- returnm ((string_append (csr_mnemonic_forwards op)
- (string_append (spc_forwards tt)
- (string_append w__158
- (string_append (sep_forwards tt)
- (string_append w__159
- (string_append (sep_forwards tt) (string_append w__160 "")))))))
- : string)
- | URET (tt) => returnm ("uret" : string)
- | ILLEGAL (s) =>
- returnm ((string_append "illegal"
- (string_append (spc_forwards tt) (string_append (decimal_string_of_bits s) "")))
- : string)
- | C_ILLEGAL (s) =>
- returnm ((string_append "c.illegal"
- (string_append (spc_forwards tt) (string_append (decimal_string_of_bits s) "")))
- : string)
+ returnm (string_append (csr_mnemonic_forwards op)
+ (string_append "i"
+ (string_append (spc_forwards tt)
+ (string_append w__156
+ (string_append (sep_forwards tt)
+ (string_append (csr_name_map_forwards csr)
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits rs1) ""))))))))
+ | CSR (csr, rs1, rd, false, op) =>
+ (reg_name_forwards rd) >>= fun w__157 : string =>
+ (reg_name_forwards rs1) >>= fun w__158 : string =>
+ returnm (string_append (csr_mnemonic_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__157
+ (string_append (sep_forwards tt)
+ (string_append (csr_name_map_forwards csr)
+ (string_append (sep_forwards tt) (string_append w__158 "")))))))
+ | URET tt => returnm "uret"
+ | LOAD_FP (imm, rs1, rd, width) =>
+ (freg_name_forwards rd) >>= fun w__159 : string =>
+ (reg_name_forwards rs1) >>= fun w__160 : string =>
+ returnm (string_append "fl"
+ (string_append (size_mnemonic_forwards width)
+ (string_append (spc_forwards tt)
+ (string_append w__159
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm)
+ (string_append (opt_spc_forwards tt)
+ (string_append "("
+ (string_append (opt_spc_forwards tt)
+ (string_append w__160
+ (string_append (opt_spc_forwards tt)
+ (string_append ")" ""))))))))))))
+ | STORE_FP (imm, rs2, rs1, width) =>
+ (freg_name_forwards rs2) >>= fun w__161 : string =>
+ (reg_name_forwards rs1) >>= fun w__162 : string =>
+ returnm (string_append "fs"
+ (string_append (size_mnemonic_forwards width)
+ (string_append (spc_forwards tt)
+ (string_append w__161
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm)
+ (string_append (opt_spc_forwards tt)
+ (string_append "("
+ (string_append (opt_spc_forwards tt)
+ (string_append w__162
+ (string_append (opt_spc_forwards tt)
+ (string_append ")" ""))))))))))))
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, op) =>
+ (freg_name_forwards rd) >>= fun w__163 : string =>
+ (freg_name_forwards rs1) >>= fun w__164 : string =>
+ (freg_name_forwards rs2) >>= fun w__165 : string =>
+ (freg_name_forwards rs3) >>= fun w__166 : string =>
+ returnm (string_append (f_madd_type_mnemonic_S_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__163
+ (string_append (sep_forwards tt)
+ (string_append w__164
+ (string_append (sep_forwards tt)
+ (string_append w__165
+ (string_append (sep_forwards tt)
+ (string_append w__166
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))))))
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, op) =>
+ (freg_name_forwards rd) >>= fun w__167 : string =>
+ (freg_name_forwards rs1) >>= fun w__168 : string =>
+ (freg_name_forwards rs2) >>= fun w__169 : string =>
+ returnm (string_append (f_bin_rm_type_mnemonic_S_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__167
+ (string_append (sep_forwards tt)
+ (string_append w__168
+ (string_append (sep_forwards tt)
+ (string_append w__169
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FSQRT_S) =>
+ (freg_name_forwards rd) >>= fun w__170 : string =>
+ (freg_name_forwards rs1) >>= fun w__171 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FSQRT_S)
+ (string_append (spc_forwards tt)
+ (string_append w__170
+ (string_append (sep_forwards tt)
+ (string_append w__171
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_W_S) =>
+ (reg_name_forwards rd) >>= fun w__172 : string =>
+ (freg_name_forwards rs1) >>= fun w__173 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FCVT_W_S)
+ (string_append (spc_forwards tt)
+ (string_append w__172
+ (string_append (sep_forwards tt)
+ (string_append w__173
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_WU_S) =>
+ (reg_name_forwards rd) >>= fun w__174 : string =>
+ (freg_name_forwards rs1) >>= fun w__175 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FCVT_WU_S)
+ (string_append (spc_forwards tt)
+ (string_append w__174
+ (string_append (sep_forwards tt)
+ (string_append w__175
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_W) =>
+ (freg_name_forwards rd) >>= fun w__176 : string =>
+ (reg_name_forwards rs1) >>= fun w__177 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FCVT_S_W)
+ (string_append (spc_forwards tt)
+ (string_append w__176
+ (string_append (sep_forwards tt)
+ (string_append w__177
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_WU) =>
+ (freg_name_forwards rd) >>= fun w__178 : string =>
+ (reg_name_forwards rs1) >>= fun w__179 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FCVT_S_WU)
+ (string_append (spc_forwards tt)
+ (string_append w__178
+ (string_append (sep_forwards tt)
+ (string_append w__179
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_L_S) =>
+ (reg_name_forwards rd) >>= fun w__180 : string =>
+ (freg_name_forwards rs1) >>= fun w__181 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FCVT_L_S)
+ (string_append (spc_forwards tt)
+ (string_append w__180
+ (string_append (sep_forwards tt)
+ (string_append w__181
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_LU_S) =>
+ (reg_name_forwards rd) >>= fun w__182 : string =>
+ (freg_name_forwards rs1) >>= fun w__183 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FCVT_LU_S)
+ (string_append (spc_forwards tt)
+ (string_append w__182
+ (string_append (sep_forwards tt)
+ (string_append w__183
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_L) =>
+ (freg_name_forwards rd) >>= fun w__184 : string =>
+ (reg_name_forwards rs1) >>= fun w__185 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FCVT_S_L)
+ (string_append (spc_forwards tt)
+ (string_append w__184
+ (string_append (sep_forwards tt)
+ (string_append w__185
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_LU) =>
+ (freg_name_forwards rd) >>= fun w__186 : string =>
+ (reg_name_forwards rs1) >>= fun w__187 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_S_forwards FCVT_S_LU)
+ (string_append (spc_forwards tt)
+ (string_append w__186
+ (string_append (sep_forwards tt)
+ (string_append w__187
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJ_S) =>
+ (freg_name_forwards rd) >>= fun w__188 : string =>
+ (freg_name_forwards rs1) >>= fun w__189 : string =>
+ (freg_name_forwards rs2) >>= fun w__190 : string =>
+ returnm (string_append (f_bin_type_mnemonic_S_forwards FSGNJ_S)
+ (string_append (spc_forwards tt)
+ (string_append w__188
+ (string_append (sep_forwards tt)
+ (string_append w__189
+ (string_append (sep_forwards tt) (string_append w__190 "")))))))
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJN_S) =>
+ (freg_name_forwards rd) >>= fun w__191 : string =>
+ (freg_name_forwards rs1) >>= fun w__192 : string =>
+ (freg_name_forwards rs2) >>= fun w__193 : string =>
+ returnm (string_append (f_bin_type_mnemonic_S_forwards FSGNJN_S)
+ (string_append (spc_forwards tt)
+ (string_append w__191
+ (string_append (sep_forwards tt)
+ (string_append w__192
+ (string_append (sep_forwards tt) (string_append w__193 "")))))))
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJX_S) =>
+ (freg_name_forwards rd) >>= fun w__194 : string =>
+ (freg_name_forwards rs1) >>= fun w__195 : string =>
+ (freg_name_forwards rs2) >>= fun w__196 : string =>
+ returnm (string_append (f_bin_type_mnemonic_S_forwards FSGNJX_S)
+ (string_append (spc_forwards tt)
+ (string_append w__194
+ (string_append (sep_forwards tt)
+ (string_append w__195
+ (string_append (sep_forwards tt) (string_append w__196 "")))))))
+ | F_BIN_TYPE_S (rs2, rs1, rd, FMIN_S) =>
+ (freg_name_forwards rd) >>= fun w__197 : string =>
+ (freg_name_forwards rs1) >>= fun w__198 : string =>
+ (freg_name_forwards rs2) >>= fun w__199 : string =>
+ returnm (string_append (f_bin_type_mnemonic_S_forwards FMIN_S)
+ (string_append (spc_forwards tt)
+ (string_append w__197
+ (string_append (sep_forwards tt)
+ (string_append w__198
+ (string_append (sep_forwards tt) (string_append w__199 "")))))))
+ | F_BIN_TYPE_S (rs2, rs1, rd, FMAX_S) =>
+ (freg_name_forwards rd) >>= fun w__200 : string =>
+ (freg_name_forwards rs1) >>= fun w__201 : string =>
+ (freg_name_forwards rs2) >>= fun w__202 : string =>
+ returnm (string_append (f_bin_type_mnemonic_S_forwards FMAX_S)
+ (string_append (spc_forwards tt)
+ (string_append w__200
+ (string_append (sep_forwards tt)
+ (string_append w__201
+ (string_append (sep_forwards tt) (string_append w__202 "")))))))
+ | F_BIN_TYPE_S (rs2, rs1, rd, FEQ_S) =>
+ (reg_name_forwards rd) >>= fun w__203 : string =>
+ (freg_name_forwards rs1) >>= fun w__204 : string =>
+ (freg_name_forwards rs2) >>= fun w__205 : string =>
+ returnm (string_append (f_bin_type_mnemonic_S_forwards FEQ_S)
+ (string_append (spc_forwards tt)
+ (string_append w__203
+ (string_append (sep_forwards tt)
+ (string_append w__204
+ (string_append (sep_forwards tt) (string_append w__205 "")))))))
+ | F_BIN_TYPE_S (rs2, rs1, rd, FLT_S) =>
+ (reg_name_forwards rd) >>= fun w__206 : string =>
+ (freg_name_forwards rs1) >>= fun w__207 : string =>
+ (freg_name_forwards rs2) >>= fun w__208 : string =>
+ returnm (string_append (f_bin_type_mnemonic_S_forwards FLT_S)
+ (string_append (spc_forwards tt)
+ (string_append w__206
+ (string_append (sep_forwards tt)
+ (string_append w__207
+ (string_append (sep_forwards tt) (string_append w__208 "")))))))
+ | F_BIN_TYPE_S (rs2, rs1, rd, FLE_S) =>
+ (reg_name_forwards rd) >>= fun w__209 : string =>
+ (freg_name_forwards rs1) >>= fun w__210 : string =>
+ (freg_name_forwards rs2) >>= fun w__211 : string =>
+ returnm (string_append (f_bin_type_mnemonic_S_forwards FLE_S)
+ (string_append (spc_forwards tt)
+ (string_append w__209
+ (string_append (sep_forwards tt)
+ (string_append w__210
+ (string_append (sep_forwards tt) (string_append w__211 "")))))))
+ | F_UN_TYPE_S (rs1, rd, FMV_X_W) =>
+ (reg_name_forwards rd) >>= fun w__212 : string =>
+ (freg_name_forwards rs1) >>= fun w__213 : string =>
+ returnm (string_append (f_un_type_mnemonic_S_forwards FMV_X_W)
+ (string_append (spc_forwards tt)
+ (string_append w__212
+ (string_append (sep_forwards tt) (string_append w__213 "")))))
+ | F_UN_TYPE_S (rs1, rd, FMV_W_X) =>
+ (freg_name_forwards rd) >>= fun w__214 : string =>
+ (reg_name_forwards rs1) >>= fun w__215 : string =>
+ returnm (string_append (f_un_type_mnemonic_S_forwards FMV_W_X)
+ (string_append (spc_forwards tt)
+ (string_append w__214
+ (string_append (sep_forwards tt) (string_append w__215 "")))))
+ | F_UN_TYPE_S (rs1, rd, FCLASS_S) =>
+ (reg_name_forwards rd) >>= fun w__216 : string =>
+ (freg_name_forwards rs1) >>= fun w__217 : string =>
+ returnm (string_append (f_un_type_mnemonic_S_forwards FCLASS_S)
+ (string_append (spc_forwards tt)
+ (string_append w__216
+ (string_append (sep_forwards tt) (string_append w__217 "")))))
+ | C_FLWSP (imm, rd) =>
+ (if sumbool_of_bool (Z.eqb 64 32) then
+ (reg_name_forwards rd) >>= fun w__218 : string =>
+ returnm (string_append "c.flwsp"
+ (string_append (spc_forwards tt)
+ (string_append w__218
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits imm) "")))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string)
+ | C_FSWSP (uimm, rd) =>
+ (if sumbool_of_bool (Z.eqb 64 32) then
+ (reg_name_forwards rd) >>= fun w__221 : string =>
+ returnm (string_append "c.fswsp"
+ (string_append (spc_forwards tt)
+ (string_append w__221
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits uimm) "")))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string)
+ | C_FLW (uimm, rsc, rdc) =>
+ (if sumbool_of_bool (Z.eqb 64 32) then
+ (creg_name_forwards rdc) >>= fun w__224 : string =>
+ (creg_name_forwards rsc) >>= fun w__225 : string =>
+ returnm (string_append "c.flw"
+ (string_append (spc_forwards tt)
+ (string_append w__224
+ (string_append (sep_forwards tt)
+ (string_append w__225
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (uimm : mword 5) ('b"00" : mword 2))) "")))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string)
+ | C_FSW (uimm, rsc1, rsc2) =>
+ (if sumbool_of_bool (Z.eqb 64 32) then
+ (creg_name_forwards rsc1) >>= fun w__228 : string =>
+ (creg_name_forwards rsc2) >>= fun w__229 : string =>
+ returnm (string_append "c.fsw"
+ (string_append (spc_forwards tt)
+ (string_append w__228
+ (string_append (sep_forwards tt)
+ (string_append w__229
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (uimm : mword 5) ('b"00" : mword 2))) "")))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string)
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, op) =>
+ (freg_name_forwards rd) >>= fun w__232 : string =>
+ (freg_name_forwards rs1) >>= fun w__233 : string =>
+ (freg_name_forwards rs2) >>= fun w__234 : string =>
+ (freg_name_forwards rs3) >>= fun w__235 : string =>
+ returnm (string_append (f_madd_type_mnemonic_D_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__232
+ (string_append (sep_forwards tt)
+ (string_append w__233
+ (string_append (sep_forwards tt)
+ (string_append w__234
+ (string_append (sep_forwards tt)
+ (string_append w__235
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))))))
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, op) =>
+ (freg_name_forwards rd) >>= fun w__236 : string =>
+ (freg_name_forwards rs1) >>= fun w__237 : string =>
+ (freg_name_forwards rs2) >>= fun w__238 : string =>
+ returnm (string_append (f_bin_rm_type_mnemonic_D_forwards op)
+ (string_append (spc_forwards tt)
+ (string_append w__236
+ (string_append (sep_forwards tt)
+ (string_append w__237
+ (string_append (sep_forwards tt)
+ (string_append w__238
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FSQRT_D) =>
+ (freg_name_forwards rd) >>= fun w__239 : string =>
+ (freg_name_forwards rs1) >>= fun w__240 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FSQRT_D)
+ (string_append (spc_forwards tt)
+ (string_append w__239
+ (string_append (sep_forwards tt)
+ (string_append w__240
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_W_D) =>
+ (reg_name_forwards rd) >>= fun w__241 : string =>
+ (freg_name_forwards rs1) >>= fun w__242 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_W_D)
+ (string_append (spc_forwards tt)
+ (string_append w__241
+ (string_append (sep_forwards tt)
+ (string_append w__242
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_WU_D) =>
+ (reg_name_forwards rd) >>= fun w__243 : string =>
+ (freg_name_forwards rs1) >>= fun w__244 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_WU_D)
+ (string_append (spc_forwards tt)
+ (string_append w__243
+ (string_append (sep_forwards tt)
+ (string_append w__244
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_W) =>
+ (freg_name_forwards rd) >>= fun w__245 : string =>
+ (reg_name_forwards rs1) >>= fun w__246 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_D_W)
+ (string_append (spc_forwards tt)
+ (string_append w__245
+ (string_append (sep_forwards tt)
+ (string_append w__246
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_WU) =>
+ (freg_name_forwards rd) >>= fun w__247 : string =>
+ (reg_name_forwards rs1) >>= fun w__248 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_D_WU)
+ (string_append (spc_forwards tt)
+ (string_append w__247
+ (string_append (sep_forwards tt)
+ (string_append w__248
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_L_D) =>
+ (reg_name_forwards rd) >>= fun w__249 : string =>
+ (freg_name_forwards rs1) >>= fun w__250 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_L_D)
+ (string_append (spc_forwards tt)
+ (string_append w__249
+ (string_append (sep_forwards tt)
+ (string_append w__250
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_LU_D) =>
+ (reg_name_forwards rd) >>= fun w__251 : string =>
+ (freg_name_forwards rs1) >>= fun w__252 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_LU_D)
+ (string_append (spc_forwards tt)
+ (string_append w__251
+ (string_append (sep_forwards tt)
+ (string_append w__252
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_L) =>
+ (freg_name_forwards rd) >>= fun w__253 : string =>
+ (reg_name_forwards rs1) >>= fun w__254 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_D_L)
+ (string_append (spc_forwards tt)
+ (string_append w__253
+ (string_append (sep_forwards tt)
+ (string_append w__254
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_LU) =>
+ (freg_name_forwards rd) >>= fun w__255 : string =>
+ (reg_name_forwards rs1) >>= fun w__256 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_D_LU)
+ (string_append (spc_forwards tt)
+ (string_append w__255
+ (string_append (sep_forwards tt)
+ (string_append w__256
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_S_D) =>
+ (freg_name_forwards rd) >>= fun w__257 : string =>
+ (freg_name_forwards rs1) >>= fun w__258 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_S_D)
+ (string_append (spc_forwards tt)
+ (string_append w__257
+ (string_append (sep_forwards tt)
+ (string_append w__258
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_S) =>
+ (freg_name_forwards rd) >>= fun w__259 : string =>
+ (freg_name_forwards rs1) >>= fun w__260 : string =>
+ returnm (string_append (f_un_rm_type_mnemonic_D_forwards FCVT_D_S)
+ (string_append (spc_forwards tt)
+ (string_append w__259
+ (string_append (sep_forwards tt)
+ (string_append w__260
+ (string_append (sep_forwards tt)
+ (string_append (frm_mnemonic_forwards rm) "")))))))
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJ_D) =>
+ (freg_name_forwards rd) >>= fun w__261 : string =>
+ (freg_name_forwards rs1) >>= fun w__262 : string =>
+ (freg_name_forwards rs2) >>= fun w__263 : string =>
+ returnm (string_append (f_bin_type_mnemonic_D_forwards FSGNJ_D)
+ (string_append (spc_forwards tt)
+ (string_append w__261
+ (string_append (sep_forwards tt)
+ (string_append w__262
+ (string_append (sep_forwards tt) (string_append w__263 "")))))))
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJN_D) =>
+ (freg_name_forwards rd) >>= fun w__264 : string =>
+ (freg_name_forwards rs1) >>= fun w__265 : string =>
+ (freg_name_forwards rs2) >>= fun w__266 : string =>
+ returnm (string_append (f_bin_type_mnemonic_D_forwards FSGNJN_D)
+ (string_append (spc_forwards tt)
+ (string_append w__264
+ (string_append (sep_forwards tt)
+ (string_append w__265
+ (string_append (sep_forwards tt) (string_append w__266 "")))))))
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJX_D) =>
+ (freg_name_forwards rd) >>= fun w__267 : string =>
+ (freg_name_forwards rs1) >>= fun w__268 : string =>
+ (freg_name_forwards rs2) >>= fun w__269 : string =>
+ returnm (string_append (f_bin_type_mnemonic_D_forwards FSGNJX_D)
+ (string_append (spc_forwards tt)
+ (string_append w__267
+ (string_append (sep_forwards tt)
+ (string_append w__268
+ (string_append (sep_forwards tt) (string_append w__269 "")))))))
+ | F_BIN_TYPE_D (rs2, rs1, rd, FMIN_D) =>
+ (freg_name_forwards rd) >>= fun w__270 : string =>
+ (freg_name_forwards rs1) >>= fun w__271 : string =>
+ (freg_name_forwards rs2) >>= fun w__272 : string =>
+ returnm (string_append (f_bin_type_mnemonic_D_forwards FMIN_D)
+ (string_append (spc_forwards tt)
+ (string_append w__270
+ (string_append (sep_forwards tt)
+ (string_append w__271
+ (string_append (sep_forwards tt) (string_append w__272 "")))))))
+ | F_BIN_TYPE_D (rs2, rs1, rd, FMAX_D) =>
+ (freg_name_forwards rd) >>= fun w__273 : string =>
+ (freg_name_forwards rs1) >>= fun w__274 : string =>
+ (freg_name_forwards rs2) >>= fun w__275 : string =>
+ returnm (string_append (f_bin_type_mnemonic_D_forwards FMAX_D)
+ (string_append (spc_forwards tt)
+ (string_append w__273
+ (string_append (sep_forwards tt)
+ (string_append w__274
+ (string_append (sep_forwards tt) (string_append w__275 "")))))))
+ | F_BIN_TYPE_D (rs2, rs1, rd, FEQ_D) =>
+ (reg_name_forwards rd) >>= fun w__276 : string =>
+ (freg_name_forwards rs1) >>= fun w__277 : string =>
+ (freg_name_forwards rs2) >>= fun w__278 : string =>
+ returnm (string_append (f_bin_type_mnemonic_D_forwards FEQ_D)
+ (string_append (spc_forwards tt)
+ (string_append w__276
+ (string_append (sep_forwards tt)
+ (string_append w__277
+ (string_append (sep_forwards tt) (string_append w__278 "")))))))
+ | F_BIN_TYPE_D (rs2, rs1, rd, FLT_D) =>
+ (reg_name_forwards rd) >>= fun w__279 : string =>
+ (freg_name_forwards rs1) >>= fun w__280 : string =>
+ (freg_name_forwards rs2) >>= fun w__281 : string =>
+ returnm (string_append (f_bin_type_mnemonic_D_forwards FLT_D)
+ (string_append (spc_forwards tt)
+ (string_append w__279
+ (string_append (sep_forwards tt)
+ (string_append w__280
+ (string_append (sep_forwards tt) (string_append w__281 "")))))))
+ | F_BIN_TYPE_D (rs2, rs1, rd, FLE_D) =>
+ (reg_name_forwards rd) >>= fun w__282 : string =>
+ (freg_name_forwards rs1) >>= fun w__283 : string =>
+ (freg_name_forwards rs2) >>= fun w__284 : string =>
+ returnm (string_append (f_bin_type_mnemonic_D_forwards FLE_D)
+ (string_append (spc_forwards tt)
+ (string_append w__282
+ (string_append (sep_forwards tt)
+ (string_append w__283
+ (string_append (sep_forwards tt) (string_append w__284 "")))))))
+ | F_UN_TYPE_D (rs1, rd, FMV_X_D) =>
+ (reg_name_forwards rd) >>= fun w__285 : string =>
+ (freg_name_forwards rs1) >>= fun w__286 : string =>
+ returnm (string_append (f_un_type_mnemonic_D_forwards FMV_X_D)
+ (string_append (spc_forwards tt)
+ (string_append w__285
+ (string_append (sep_forwards tt) (string_append w__286 "")))))
+ | F_UN_TYPE_D (rs1, rd, FMV_D_X) =>
+ (freg_name_forwards rd) >>= fun w__287 : string =>
+ (reg_name_forwards rs1) >>= fun w__288 : string =>
+ returnm (string_append (f_un_type_mnemonic_D_forwards FMV_D_X)
+ (string_append (spc_forwards tt)
+ (string_append w__287
+ (string_append (sep_forwards tt) (string_append w__288 "")))))
+ | F_UN_TYPE_D (rs1, rd, FCLASS_D) =>
+ (reg_name_forwards rd) >>= fun w__289 : string =>
+ (freg_name_forwards rs1) >>= fun w__290 : string =>
+ returnm (string_append (f_un_type_mnemonic_D_forwards FCLASS_D)
+ (string_append (spc_forwards tt)
+ (string_append w__289
+ (string_append (sep_forwards tt) (string_append w__290 "")))))
+ | C_FLDSP (uimm, rd) =>
+ (if sumbool_of_bool (orb (Z.eqb 64 32) (Z.eqb 64 64)) then
+ (reg_name_forwards rd) >>= fun w__291 : string =>
+ returnm (string_append "c.fldsp"
+ (string_append (spc_forwards tt)
+ (string_append w__291
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits uimm) "")))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string)
+ | C_FSDSP (uimm, rs2) =>
+ (if sumbool_of_bool (orb (Z.eqb 64 32) (Z.eqb 64 64)) then
+ (reg_name_forwards rs2) >>= fun w__294 : string =>
+ returnm (string_append "c.fsdsp"
+ (string_append (spc_forwards tt)
+ (string_append w__294
+ (string_append (sep_forwards tt)
+ (string_append (decimal_string_of_bits uimm) "")))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string)
+ | C_FLD (uimm, rsc, rdc) =>
+ (if sumbool_of_bool (orb (Z.eqb 64 32) (Z.eqb 64 64)) then
+ (creg_name_forwards rdc) >>= fun w__297 : string =>
+ (creg_name_forwards rsc) >>= fun w__298 : string =>
+ returnm (string_append "c.fld"
+ (string_append (spc_forwards tt)
+ (string_append w__297
+ (string_append (sep_forwards tt)
+ (string_append w__298
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (uimm : mword 5) ('b"000" : mword 3))) "")))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string)
+ | C_FSD (uimm, rsc1, rsc2) =>
+ (if sumbool_of_bool (orb (Z.eqb 64 32) (Z.eqb 64 64)) then
+ (creg_name_forwards rsc1) >>= fun w__301 : string =>
+ (creg_name_forwards rsc2) >>= fun w__302 : string =>
+ returnm (string_append "c.fsd"
+ (string_append (spc_forwards tt)
+ (string_append w__301
+ (string_append (sep_forwards tt)
+ (string_append w__302
+ (string_append (sep_forwards tt)
+ (string_append
+ (decimal_string_of_bits
+ (concat_vec (uimm : mword 5) ('b"000" : mword 3))) "")))))))
+ else assert_exp' false "Pattern match failure at unknown location" >>= fun _ => exit tt)
+ : M (string)
+ | ILLEGAL s =>
+ returnm (string_append "illegal"
+ (string_append (spc_forwards tt) (string_append (decimal_string_of_bits s) "")))
+ | C_ILLEGAL s =>
+ returnm (string_append "c.illegal"
+ (string_append (spc_forwards tt) (string_append (decimal_string_of_bits s) "")))
end)
: M (string).
-Definition _s1677_ (_s1678_ : string)
-: M (option (mword 16)) :=
-
- let _s1679_ := _s1678_ in
- (if ((string_startswith _s1679_ "c.illegal")) then
- (match (string_drop _s1679_ (projT1 (string_length "c.illegal"))) with
- | _s1680_ =>
- (spc_matches_prefix _s1680_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s1681_ _)) =>
- match (string_drop _s1680_ _s1681_) with
- | _s1682_ =>
- match (hex_bits_16_matches_prefix _s1682_) with
- | Some ((s, existT _ _s1683_ _)) =>
- let p0_ := string_drop _s1682_ _s1683_ in
- if ((generic_eq p0_ "")) then Some (s)
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 16))
+Definition _s3138_ (_s3139_ : string) : M (option (mword 16)) :=
+ let _s3140_ := _s3139_ in
+ (if string_startswith _s3140_ "c.illegal" then
+ (match (string_drop _s3140_ (projT1 (string_length "c.illegal"))) with
+ | _s3141_ =>
+ (spc_matches_prefix _s3141_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s3142_ _) =>
+ match (string_drop _s3141_ _s3142_) with
+ | _s3143_ =>
+ match (hex_bits_16_matches_prefix _s3143_) with
+ | Some (s, existT _ _s3144_ _) =>
+ let p0_ := string_drop _s3143_ _s3144_ in
+ if generic_eq p0_ "" then Some s
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 16))
- else returnm (None : option (mword 16)))
+ else returnm None)
: M (option (mword 16)).
-Definition _s1669_ (_s1670_ : string)
-: M (option (mword 32)) :=
-
- let _s1671_ := _s1670_ in
- (if ((string_startswith _s1671_ "illegal")) then
- (match (string_drop _s1671_ (projT1 (string_length "illegal"))) with
- | _s1672_ =>
- (spc_matches_prefix _s1672_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s1673_ _)) =>
- match (string_drop _s1672_ _s1673_) with
- | _s1674_ =>
- match (hex_bits_32_matches_prefix _s1674_) with
- | Some ((s, existT _ _s1675_ _)) =>
- let p0_ := string_drop _s1674_ _s1675_ in
- if ((generic_eq p0_ "")) then Some (s)
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 32))
+Definition _s3130_ (_s3131_ : string) : M (option (mword 32)) :=
+ let _s3132_ := _s3131_ in
+ (if string_startswith _s3132_ "illegal" then
+ (match (string_drop _s3132_ (projT1 (string_length "illegal"))) with
+ | _s3133_ =>
+ (spc_matches_prefix _s3133_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s3134_ _) =>
+ match (string_drop _s3133_ _s3134_) with
+ | _s3135_ =>
+ match (hex_bits_32_matches_prefix _s3135_) with
+ | Some (s, existT _ _s3136_ _) =>
+ let p0_ := string_drop _s3135_ _s3136_ in
+ if generic_eq p0_ "" then Some s
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 32))
- else returnm (None : option (mword 32)))
+ else returnm None)
: M (option (mword 32)).
-Definition _s1652_ (_s1653_ : string)
-: M (option ((csrop * mword 5 * mword 5 * mword 12))) :=
-
- (match _s1653_ with
- | _s1654_ =>
- (csr_mnemonic_matches_prefix _s1654_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=
+Definition _s3114_ (_s3115_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s3116_ := _s3115_ in
+ (if string_startswith _s3116_ "c.fsd" then
+ (match (string_drop _s3116_ (projT1 (string_length "c.fsd"))) with
+ | _s3117_ =>
+ (spc_matches_prefix _s3117_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s3118_ _) =>
+ (match (string_drop _s3117_ _s3118_) with
+ | _s3119_ =>
+ (creg_name_matches_prefix _s3119_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rsc1, existT _ _s3120_ _) =>
+ (match (string_drop _s3119_ _s3120_) with
+ | _s3121_ =>
+ (sep_matches_prefix _s3121_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s3122_ _) =>
+ (match (string_drop _s3121_ _s3122_) with
+ | _s3123_ =>
+ (creg_name_matches_prefix _s3123_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc2, existT _ _s3124_ _) =>
+ (match (string_drop _s3123_ _s3124_) with
+ | _s3125_ =>
+ (sep_matches_prefix _s3125_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s3126_ _) =>
+ match (string_drop _s3125_ _s3126_) with
+ | _s3127_ =>
+ match (hex_bits_8_matches_prefix _s3127_) with
+ | Some (v__1338, existT _ _s3128_ _) =>
+ if eq_vec (subrange_vec_dec v__1338 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1338 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1338 7 3 in
+ let p0_ := string_drop _s3127_ _s3128_ in
+ if generic_eq p0_ "" then
+ Some (rsc1, rsc2, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5))).
+
+Definition _s3098_ (_s3099_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s3100_ := _s3099_ in
+ (if string_startswith _s3100_ "c.fld" then
+ (match (string_drop _s3100_ (projT1 (string_length "c.fld"))) with
+ | _s3101_ =>
+ (spc_matches_prefix _s3101_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s3102_ _) =>
+ (match (string_drop _s3101_ _s3102_) with
+ | _s3103_ =>
+ (creg_name_matches_prefix _s3103_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rdc, existT _ _s3104_ _) =>
+ (match (string_drop _s3103_ _s3104_) with
+ | _s3105_ =>
+ (sep_matches_prefix _s3105_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s3106_ _) =>
+ (match (string_drop _s3105_ _s3106_) with
+ | _s3107_ =>
+ (creg_name_matches_prefix _s3107_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc, existT _ _s3108_ _) =>
+ (match (string_drop _s3107_ _s3108_) with
+ | _s3109_ =>
+ (sep_matches_prefix _s3109_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s3110_ _) =>
+ match (string_drop _s3109_ _s3110_) with
+ | _s3111_ =>
+ match (hex_bits_8_matches_prefix _s3111_) with
+ | Some (v__1340, existT _ _s3112_ _) =>
+ if eq_vec (subrange_vec_dec v__1340 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1340 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1340 7 3 in
+ let p0_ := string_drop _s3111_ _s3112_ in
+ if generic_eq p0_ "" then
+ Some (rdc, rsc, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5))).
+
+Definition _s3086_ (_s3087_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3088_ := _s3087_ in
+ (if string_startswith _s3088_ "c.fsdsp" then
+ (match (string_drop _s3088_ (projT1 (string_length "c.fsdsp"))) with
+ | _s3089_ =>
+ (spc_matches_prefix _s3089_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s3090_ _) =>
+ (match (string_drop _s3089_ _s3090_) with
+ | _s3091_ =>
+ (reg_name_matches_prefix _s3091_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rs2, existT _ _s3092_ _) =>
+ (match (string_drop _s3091_ _s3092_) with
+ | _s3093_ =>
+ (sep_matches_prefix _s3093_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3094_ _) =>
+ match (string_drop _s3093_ _s3094_) with
+ | _s3095_ =>
+ match (hex_bits_6_matches_prefix _s3095_) with
+ | Some (uimm, existT _ _s3096_ _) =>
+ let p0_ := string_drop _s3095_ _s3096_ in
+ if generic_eq p0_ "" then Some (rs2, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6))).
+
+Definition _s3074_ (_s3075_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3076_ := _s3075_ in
+ (if string_startswith _s3076_ "c.fldsp" then
+ (match (string_drop _s3076_ (projT1 (string_length "c.fldsp"))) with
+ | _s3077_ =>
+ (spc_matches_prefix _s3077_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s3078_ _) =>
+ (match (string_drop _s3077_ _s3078_) with
+ | _s3079_ =>
+ (reg_name_matches_prefix _s3079_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s3080_ _) =>
+ (match (string_drop _s3079_ _s3080_) with
+ | _s3081_ =>
+ (sep_matches_prefix _s3081_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3082_ _) =>
+ match (string_drop _s3081_ _s3082_) with
+ | _s3083_ =>
+ match (hex_bits_6_matches_prefix _s3083_) with
+ | Some (uimm, existT _ _s3084_ _) =>
+ let p0_ := string_drop _s3083_ _s3084_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6))).
+
+Definition _s3061_ (_s3062_ : string) : M (option ((f_un_op_D * mword 5 * mword 5))) :=
+ (match _s3062_ with
+ | _s3063_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s3063_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCLASS_D, existT _ _s3064_ _) =>
+ (match (string_drop _s3063_ _s3064_) with
+ | _s3065_ =>
+ (spc_matches_prefix _s3065_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s3066_ _) =>
+ (match (string_drop _s3065_ _s3066_) with
+ | _s3067_ =>
+ (reg_name_matches_prefix _s3067_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s3068_ _) =>
+ (match (string_drop _s3067_ _s3068_) with
+ | _s3069_ =>
+ (sep_matches_prefix _s3069_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s3070_ _) =>
+ (match (string_drop _s3069_ _s3070_) with
+ | _s3071_ =>
+ (freg_name_matches_prefix _s3071_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s3072_ _) =>
+ let p0_ := string_drop _s3071_ _s3072_ in
+ if generic_eq p0_ "" then
+ Some (FCLASS_D, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5))).
+
+Definition _s3048_ (_s3049_ : string) : M (option ((f_un_op_D * mword 5 * mword 5))) :=
+ (match _s3049_ with
+ | _s3050_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s3050_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1655_ _)) =>
- (match (string_drop _s1654_ _s1655_) with
- | _s1656_ =>
- (spc_matches_prefix _s1656_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (FMV_D_X, existT _ _s3051_ _) =>
+ (match (string_drop _s3050_ _s3051_) with
+ | _s3052_ =>
+ (spc_matches_prefix _s3052_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1657_ _)) =>
- (match (string_drop _s1656_ _s1657_) with
- | _s1658_ =>
- (reg_name_matches_prefix _s1658_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3053_ _) =>
+ (match (string_drop _s3052_ _s3053_) with
+ | _s3054_ =>
+ (freg_name_matches_prefix _s3054_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1659_ _)) =>
- (match (string_drop _s1658_ _s1659_) with
- | _s1660_ =>
- (sep_matches_prefix _s1660_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3055_ _) =>
+ (match (string_drop _s3054_ _s3055_) with
+ | _s3056_ =>
+ (sep_matches_prefix _s3056_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1661_ _)) =>
- (match (string_drop _s1660_ _s1661_) with
- | _s1662_ =>
- (reg_name_matches_prefix _s1662_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3057_ _) =>
+ (match (string_drop _s3056_ _s3057_) with
+ | _s3058_ =>
+ (reg_name_matches_prefix _s3058_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s3059_ _) =>
+ let p0_ := string_drop _s3058_ _s3059_ in
+ if generic_eq p0_ "" then Some (FMV_D_X, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5))).
+
+Definition _s3035_ (_s3036_ : string) : M (option ((f_un_op_D * mword 5 * mword 5))) :=
+ (match _s3036_ with
+ | _s3037_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s3037_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_X_D, existT _ _s3038_ _) =>
+ (match (string_drop _s3037_ _s3038_) with
+ | _s3039_ =>
+ (spc_matches_prefix _s3039_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s3040_ _) =>
+ (match (string_drop _s3039_ _s3040_) with
+ | _s3041_ =>
+ (reg_name_matches_prefix _s3041_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s3042_ _) =>
+ (match (string_drop _s3041_ _s3042_) with
+ | _s3043_ =>
+ (sep_matches_prefix _s3043_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s3044_ _) =>
+ (match (string_drop _s3043_ _s3044_) with
+ | _s3045_ =>
+ (freg_name_matches_prefix _s3045_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s3046_ _) =>
+ let p0_ := string_drop _s3045_ _s3046_ in
+ if generic_eq p0_ "" then Some (FMV_X_D, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5))).
+
+Definition _s3018_ (_s3019_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s3019_ with
+ | _s3020_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s3020_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLE_D, existT _ _s3021_ _) =>
+ (match (string_drop _s3020_ _s3021_) with
+ | _s3022_ =>
+ (spc_matches_prefix _s3022_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s3023_ _) =>
+ (match (string_drop _s3022_ _s3023_) with
+ | _s3024_ =>
+ (reg_name_matches_prefix _s3024_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s3025_ _) =>
+ (match (string_drop _s3024_ _s3025_) with
+ | _s3026_ =>
+ (sep_matches_prefix _s3026_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s3027_ _) =>
+ (match (string_drop _s3026_ _s3027_) with
+ | _s3028_ =>
+ (freg_name_matches_prefix _s3028_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1663_ _)) =>
- (match (string_drop _s1662_ _s1663_) with
- | _s1664_ =>
- (sep_matches_prefix _s1664_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3029_ _) =>
+ (match (string_drop _s3028_ _s3029_) with
+ | _s3030_ =>
+ (sep_matches_prefix _s3030_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1665_ _)) =>
- (match (string_drop _s1664_ _s1665_) with
- | _s1666_ =>
- (csr_name_map_matches_prefix _s1666_) >>= fun w__6 : option ((mword 12 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3031_ _) =>
+ (match (string_drop _s3030_ _s3031_) with
+ | _s3032_ =>
+ (freg_name_matches_prefix _s3032_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((csr, existT _ _s1667_ _)) =>
- let p0_ :=
- string_drop _s1666_ _s1667_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, csr))
- else None
- | _ => None
- end)
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s3033_ _) =>
+ let p0_ :=
+ string_drop _s3032_ _s3033_ in
+ if generic_eq p0_ "" then
+ Some (FLE_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s3001_ (_s3002_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s3002_ with
+ | _s3003_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s3003_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLT_D, existT _ _s3004_ _) =>
+ (match (string_drop _s3003_ _s3004_) with
+ | _s3005_ =>
+ (spc_matches_prefix _s3005_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s3006_ _) =>
+ (match (string_drop _s3005_ _s3006_) with
+ | _s3007_ =>
+ (reg_name_matches_prefix _s3007_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s3008_ _) =>
+ (match (string_drop _s3007_ _s3008_) with
+ | _s3009_ =>
+ (sep_matches_prefix _s3009_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s3010_ _) =>
+ (match (string_drop _s3009_ _s3010_) with
+ | _s3011_ =>
+ (freg_name_matches_prefix _s3011_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s3012_ _) =>
+ (match (string_drop _s3011_ _s3012_) with
+ | _s3013_ =>
+ (sep_matches_prefix _s3013_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s3014_ _) =>
+ (match (string_drop _s3013_ _s3014_) with
+ | _s3015_ =>
+ (freg_name_matches_prefix _s3015_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s3016_ _) =>
+ let p0_ :=
+ string_drop _s3015_ _s3016_ in
+ if generic_eq p0_ "" then
+ Some (FLT_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s2984_ (_s2985_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s2985_ with
+ | _s2986_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s2986_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FEQ_D, existT _ _s2987_ _) =>
+ (match (string_drop _s2986_ _s2987_) with
+ | _s2988_ =>
+ (spc_matches_prefix _s2988_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2989_ _) =>
+ (match (string_drop _s2988_ _s2989_) with
+ | _s2990_ =>
+ (reg_name_matches_prefix _s2990_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2991_ _) =>
+ (match (string_drop _s2990_ _s2991_) with
+ | _s2992_ =>
+ (sep_matches_prefix _s2992_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2993_ _) =>
+ (match (string_drop _s2992_ _s2993_) with
+ | _s2994_ =>
+ (freg_name_matches_prefix _s2994_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2995_ _) =>
+ (match (string_drop _s2994_ _s2995_) with
+ | _s2996_ =>
+ (sep_matches_prefix _s2996_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2997_ _) =>
+ (match (string_drop _s2996_ _s2997_) with
+ | _s2998_ =>
+ (freg_name_matches_prefix _s2998_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2999_ _) =>
+ let p0_ :=
+ string_drop _s2998_ _s2999_ in
+ if generic_eq p0_ "" then
+ Some (FEQ_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s2967_ (_s2968_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s2968_ with
+ | _s2969_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s2969_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMAX_D, existT _ _s2970_ _) =>
+ (match (string_drop _s2969_ _s2970_) with
+ | _s2971_ =>
+ (spc_matches_prefix _s2971_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2972_ _) =>
+ (match (string_drop _s2971_ _s2972_) with
+ | _s2973_ =>
+ (freg_name_matches_prefix _s2973_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2974_ _) =>
+ (match (string_drop _s2973_ _s2974_) with
+ | _s2975_ =>
+ (sep_matches_prefix _s2975_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2976_ _) =>
+ (match (string_drop _s2975_ _s2976_) with
+ | _s2977_ =>
+ (freg_name_matches_prefix _s2977_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2978_ _) =>
+ (match (string_drop _s2977_ _s2978_) with
+ | _s2979_ =>
+ (sep_matches_prefix _s2979_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2980_ _) =>
+ (match (string_drop _s2979_ _s2980_) with
+ | _s2981_ =>
+ (freg_name_matches_prefix _s2981_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2982_ _) =>
+ let p0_ :=
+ string_drop _s2981_ _s2982_ in
+ if generic_eq p0_ "" then
+ Some (FMAX_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s2950_ (_s2951_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s2951_ with
+ | _s2952_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s2952_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMIN_D, existT _ _s2953_ _) =>
+ (match (string_drop _s2952_ _s2953_) with
+ | _s2954_ =>
+ (spc_matches_prefix _s2954_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2955_ _) =>
+ (match (string_drop _s2954_ _s2955_) with
+ | _s2956_ =>
+ (freg_name_matches_prefix _s2956_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2957_ _) =>
+ (match (string_drop _s2956_ _s2957_) with
+ | _s2958_ =>
+ (sep_matches_prefix _s2958_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2959_ _) =>
+ (match (string_drop _s2958_ _s2959_) with
+ | _s2960_ =>
+ (freg_name_matches_prefix _s2960_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2961_ _) =>
+ (match (string_drop _s2960_ _s2961_) with
+ | _s2962_ =>
+ (sep_matches_prefix _s2962_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2963_ _) =>
+ (match (string_drop _s2962_ _s2963_) with
+ | _s2964_ =>
+ (freg_name_matches_prefix _s2964_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2965_ _) =>
+ let p0_ :=
+ string_drop _s2964_ _s2965_ in
+ if generic_eq p0_ "" then
+ Some (FMIN_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s2933_ (_s2934_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s2934_ with
+ | _s2935_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s2935_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJX_D, existT _ _s2936_ _) =>
+ (match (string_drop _s2935_ _s2936_) with
+ | _s2937_ =>
+ (spc_matches_prefix _s2937_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2938_ _) =>
+ (match (string_drop _s2937_ _s2938_) with
+ | _s2939_ =>
+ (freg_name_matches_prefix _s2939_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2940_ _) =>
+ (match (string_drop _s2939_ _s2940_) with
+ | _s2941_ =>
+ (sep_matches_prefix _s2941_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2942_ _) =>
+ (match (string_drop _s2941_ _s2942_) with
+ | _s2943_ =>
+ (freg_name_matches_prefix _s2943_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2944_ _) =>
+ (match (string_drop _s2943_ _s2944_) with
+ | _s2945_ =>
+ (sep_matches_prefix _s2945_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2946_ _) =>
+ (match (string_drop _s2945_ _s2946_) with
+ | _s2947_ =>
+ (freg_name_matches_prefix _s2947_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2948_ _) =>
+ let p0_ :=
+ string_drop _s2947_ _s2948_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJX_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s2916_ (_s2917_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s2917_ with
+ | _s2918_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s2918_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJN_D, existT _ _s2919_ _) =>
+ (match (string_drop _s2918_ _s2919_) with
+ | _s2920_ =>
+ (spc_matches_prefix _s2920_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2921_ _) =>
+ (match (string_drop _s2920_ _s2921_) with
+ | _s2922_ =>
+ (freg_name_matches_prefix _s2922_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2923_ _) =>
+ (match (string_drop _s2922_ _s2923_) with
+ | _s2924_ =>
+ (sep_matches_prefix _s2924_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2925_ _) =>
+ (match (string_drop _s2924_ _s2925_) with
+ | _s2926_ =>
+ (freg_name_matches_prefix _s2926_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2927_ _) =>
+ (match (string_drop _s2926_ _s2927_) with
+ | _s2928_ =>
+ (sep_matches_prefix _s2928_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2929_ _) =>
+ (match (string_drop _s2928_ _s2929_) with
+ | _s2930_ =>
+ (freg_name_matches_prefix _s2930_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2931_ _) =>
+ let p0_ :=
+ string_drop _s2930_ _s2931_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJN_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s2899_ (_s2900_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s2900_ with
+ | _s2901_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s2901_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJ_D, existT _ _s2902_ _) =>
+ (match (string_drop _s2901_ _s2902_) with
+ | _s2903_ =>
+ (spc_matches_prefix _s2903_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2904_ _) =>
+ (match (string_drop _s2903_ _s2904_) with
+ | _s2905_ =>
+ (freg_name_matches_prefix _s2905_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2906_ _) =>
+ (match (string_drop _s2905_ _s2906_) with
+ | _s2907_ =>
+ (sep_matches_prefix _s2907_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2908_ _) =>
+ (match (string_drop _s2907_ _s2908_) with
+ | _s2909_ =>
+ (freg_name_matches_prefix _s2909_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2910_ _) =>
+ (match (string_drop _s2909_ _s2910_) with
+ | _s2911_ =>
+ (sep_matches_prefix _s2911_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2912_ _) =>
+ (match (string_drop _s2911_ _s2912_) with
+ | _s2913_ =>
+ (freg_name_matches_prefix _s2913_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2914_ _) =>
+ let p0_ :=
+ string_drop _s2913_ _s2914_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJ_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s2882_ (_s2883_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2883_ with
+ | _s2884_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2884_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_S, existT _ _s2885_ _) =>
+ (match (string_drop _s2884_ _s2885_) with
+ | _s2886_ =>
+ (spc_matches_prefix _s2886_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2887_ _) =>
+ (match (string_drop _s2886_ _s2887_) with
+ | _s2888_ =>
+ (freg_name_matches_prefix _s2888_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2889_ _) =>
+ (match (string_drop _s2888_ _s2889_) with
+ | _s2890_ =>
+ (sep_matches_prefix _s2890_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2891_ _) =>
+ (match (string_drop _s2890_ _s2891_) with
+ | _s2892_ =>
+ (freg_name_matches_prefix _s2892_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2893_ _) =>
+ (match (string_drop _s2892_ _s2893_) with
+ | _s2894_ =>
+ (sep_matches_prefix _s2894_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2895_ _) =>
+ (match (string_drop _s2894_ _s2895_) with
+ | _s2896_ =>
+ (frm_mnemonic_matches_prefix _s2896_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2897_ _) =>
+ let p0_ :=
+ string_drop _s2896_ _s2897_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2865_ (_s2866_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2866_ with
+ | _s2867_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2867_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_D, existT _ _s2868_ _) =>
+ (match (string_drop _s2867_ _s2868_) with
+ | _s2869_ =>
+ (spc_matches_prefix _s2869_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2870_ _) =>
+ (match (string_drop _s2869_ _s2870_) with
+ | _s2871_ =>
+ (freg_name_matches_prefix _s2871_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2872_ _) =>
+ (match (string_drop _s2871_ _s2872_) with
+ | _s2873_ =>
+ (sep_matches_prefix _s2873_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2874_ _) =>
+ (match (string_drop _s2873_ _s2874_) with
+ | _s2875_ =>
+ (freg_name_matches_prefix _s2875_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2876_ _) =>
+ (match (string_drop _s2875_ _s2876_) with
+ | _s2877_ =>
+ (sep_matches_prefix _s2877_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2878_ _) =>
+ (match (string_drop _s2877_ _s2878_) with
+ | _s2879_ =>
+ (frm_mnemonic_matches_prefix _s2879_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2880_ _) =>
+ let p0_ :=
+ string_drop _s2879_ _s2880_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2848_ (_s2849_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2849_ with
+ | _s2850_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2850_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_LU, existT _ _s2851_ _) =>
+ (match (string_drop _s2850_ _s2851_) with
+ | _s2852_ =>
+ (spc_matches_prefix _s2852_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2853_ _) =>
+ (match (string_drop _s2852_ _s2853_) with
+ | _s2854_ =>
+ (freg_name_matches_prefix _s2854_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2855_ _) =>
+ (match (string_drop _s2854_ _s2855_) with
+ | _s2856_ =>
+ (sep_matches_prefix _s2856_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2857_ _) =>
+ (match (string_drop _s2856_ _s2857_) with
+ | _s2858_ =>
+ (reg_name_matches_prefix _s2858_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2859_ _) =>
+ (match (string_drop _s2858_ _s2859_) with
+ | _s2860_ =>
+ (sep_matches_prefix _s2860_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2861_ _) =>
+ (match (string_drop _s2860_ _s2861_) with
+ | _s2862_ =>
+ (frm_mnemonic_matches_prefix _s2862_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2863_ _) =>
+ let p0_ :=
+ string_drop _s2862_ _s2863_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_LU, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2831_ (_s2832_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2832_ with
+ | _s2833_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2833_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_L, existT _ _s2834_ _) =>
+ (match (string_drop _s2833_ _s2834_) with
+ | _s2835_ =>
+ (spc_matches_prefix _s2835_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2836_ _) =>
+ (match (string_drop _s2835_ _s2836_) with
+ | _s2837_ =>
+ (freg_name_matches_prefix _s2837_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2838_ _) =>
+ (match (string_drop _s2837_ _s2838_) with
+ | _s2839_ =>
+ (sep_matches_prefix _s2839_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2840_ _) =>
+ (match (string_drop _s2839_ _s2840_) with
+ | _s2841_ =>
+ (reg_name_matches_prefix _s2841_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2842_ _) =>
+ (match (string_drop _s2841_ _s2842_) with
+ | _s2843_ =>
+ (sep_matches_prefix _s2843_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2844_ _) =>
+ (match (string_drop _s2843_ _s2844_) with
+ | _s2845_ =>
+ (frm_mnemonic_matches_prefix _s2845_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2846_ _) =>
+ let p0_ :=
+ string_drop _s2845_ _s2846_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_L, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2814_ (_s2815_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2815_ with
+ | _s2816_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2816_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_LU_D, existT _ _s2817_ _) =>
+ (match (string_drop _s2816_ _s2817_) with
+ | _s2818_ =>
+ (spc_matches_prefix _s2818_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2819_ _) =>
+ (match (string_drop _s2818_ _s2819_) with
+ | _s2820_ =>
+ (reg_name_matches_prefix _s2820_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2821_ _) =>
+ (match (string_drop _s2820_ _s2821_) with
+ | _s2822_ =>
+ (sep_matches_prefix _s2822_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2823_ _) =>
+ (match (string_drop _s2822_ _s2823_) with
+ | _s2824_ =>
+ (freg_name_matches_prefix _s2824_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2825_ _) =>
+ (match (string_drop _s2824_ _s2825_) with
+ | _s2826_ =>
+ (sep_matches_prefix _s2826_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2827_ _) =>
+ (match (string_drop _s2826_ _s2827_) with
+ | _s2828_ =>
+ (frm_mnemonic_matches_prefix _s2828_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2829_ _) =>
+ let p0_ :=
+ string_drop _s2828_ _s2829_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_LU_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2797_ (_s2798_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2798_ with
+ | _s2799_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2799_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_L_D, existT _ _s2800_ _) =>
+ (match (string_drop _s2799_ _s2800_) with
+ | _s2801_ =>
+ (spc_matches_prefix _s2801_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2802_ _) =>
+ (match (string_drop _s2801_ _s2802_) with
+ | _s2803_ =>
+ (reg_name_matches_prefix _s2803_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2804_ _) =>
+ (match (string_drop _s2803_ _s2804_) with
+ | _s2805_ =>
+ (sep_matches_prefix _s2805_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2806_ _) =>
+ (match (string_drop _s2805_ _s2806_) with
+ | _s2807_ =>
+ (freg_name_matches_prefix _s2807_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2808_ _) =>
+ (match (string_drop _s2807_ _s2808_) with
+ | _s2809_ =>
+ (sep_matches_prefix _s2809_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2810_ _) =>
+ (match (string_drop _s2809_ _s2810_) with
+ | _s2811_ =>
+ (frm_mnemonic_matches_prefix _s2811_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2812_ _) =>
+ let p0_ :=
+ string_drop _s2811_ _s2812_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_L_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2780_ (_s2781_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2781_ with
+ | _s2782_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2782_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_WU, existT _ _s2783_ _) =>
+ (match (string_drop _s2782_ _s2783_) with
+ | _s2784_ =>
+ (spc_matches_prefix _s2784_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2785_ _) =>
+ (match (string_drop _s2784_ _s2785_) with
+ | _s2786_ =>
+ (freg_name_matches_prefix _s2786_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2787_ _) =>
+ (match (string_drop _s2786_ _s2787_) with
+ | _s2788_ =>
+ (sep_matches_prefix _s2788_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2789_ _) =>
+ (match (string_drop _s2788_ _s2789_) with
+ | _s2790_ =>
+ (reg_name_matches_prefix _s2790_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2791_ _) =>
+ (match (string_drop _s2790_ _s2791_) with
+ | _s2792_ =>
+ (sep_matches_prefix _s2792_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2793_ _) =>
+ (match (string_drop _s2792_ _s2793_) with
+ | _s2794_ =>
+ (frm_mnemonic_matches_prefix _s2794_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2795_ _) =>
+ let p0_ :=
+ string_drop _s2794_ _s2795_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_WU, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2763_ (_s2764_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2764_ with
+ | _s2765_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2765_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_W, existT _ _s2766_ _) =>
+ (match (string_drop _s2765_ _s2766_) with
+ | _s2767_ =>
+ (spc_matches_prefix _s2767_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2768_ _) =>
+ (match (string_drop _s2767_ _s2768_) with
+ | _s2769_ =>
+ (freg_name_matches_prefix _s2769_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2770_ _) =>
+ (match (string_drop _s2769_ _s2770_) with
+ | _s2771_ =>
+ (sep_matches_prefix _s2771_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2772_ _) =>
+ (match (string_drop _s2771_ _s2772_) with
+ | _s2773_ =>
+ (reg_name_matches_prefix _s2773_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2774_ _) =>
+ (match (string_drop _s2773_ _s2774_) with
+ | _s2775_ =>
+ (sep_matches_prefix _s2775_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2776_ _) =>
+ (match (string_drop _s2775_ _s2776_) with
+ | _s2777_ =>
+ (frm_mnemonic_matches_prefix _s2777_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2778_ _) =>
+ let p0_ :=
+ string_drop _s2777_ _s2778_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_W, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2746_ (_s2747_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2747_ with
+ | _s2748_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2748_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_WU_D, existT _ _s2749_ _) =>
+ (match (string_drop _s2748_ _s2749_) with
+ | _s2750_ =>
+ (spc_matches_prefix _s2750_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2751_ _) =>
+ (match (string_drop _s2750_ _s2751_) with
+ | _s2752_ =>
+ (reg_name_matches_prefix _s2752_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2753_ _) =>
+ (match (string_drop _s2752_ _s2753_) with
+ | _s2754_ =>
+ (sep_matches_prefix _s2754_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2755_ _) =>
+ (match (string_drop _s2754_ _s2755_) with
+ | _s2756_ =>
+ (freg_name_matches_prefix _s2756_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2757_ _) =>
+ (match (string_drop _s2756_ _s2757_) with
+ | _s2758_ =>
+ (sep_matches_prefix _s2758_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2759_ _) =>
+ (match (string_drop _s2758_ _s2759_) with
+ | _s2760_ =>
+ (frm_mnemonic_matches_prefix _s2760_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2761_ _) =>
+ let p0_ :=
+ string_drop _s2760_ _s2761_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_WU_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2729_ (_s2730_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2730_ with
+ | _s2731_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2731_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_W_D, existT _ _s2732_ _) =>
+ (match (string_drop _s2731_ _s2732_) with
+ | _s2733_ =>
+ (spc_matches_prefix _s2733_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2734_ _) =>
+ (match (string_drop _s2733_ _s2734_) with
+ | _s2735_ =>
+ (reg_name_matches_prefix _s2735_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2736_ _) =>
+ (match (string_drop _s2735_ _s2736_) with
+ | _s2737_ =>
+ (sep_matches_prefix _s2737_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2738_ _) =>
+ (match (string_drop _s2737_ _s2738_) with
+ | _s2739_ =>
+ (freg_name_matches_prefix _s2739_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2740_ _) =>
+ (match (string_drop _s2739_ _s2740_) with
+ | _s2741_ =>
+ (sep_matches_prefix _s2741_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2742_ _) =>
+ (match (string_drop _s2741_ _s2742_) with
+ | _s2743_ =>
+ (frm_mnemonic_matches_prefix _s2743_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2744_ _) =>
+ let p0_ :=
+ string_drop _s2743_ _s2744_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_W_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2712_ (_s2713_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2713_ with
+ | _s2714_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s2714_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSQRT_D, existT _ _s2715_ _) =>
+ (match (string_drop _s2714_ _s2715_) with
+ | _s2716_ =>
+ (spc_matches_prefix _s2716_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2717_ _) =>
+ (match (string_drop _s2716_ _s2717_) with
+ | _s2718_ =>
+ (freg_name_matches_prefix _s2718_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2719_ _) =>
+ (match (string_drop _s2718_ _s2719_) with
+ | _s2720_ =>
+ (sep_matches_prefix _s2720_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2721_ _) =>
+ (match (string_drop _s2720_ _s2721_) with
+ | _s2722_ =>
+ (freg_name_matches_prefix _s2722_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2723_ _) =>
+ (match (string_drop _s2722_ _s2723_) with
+ | _s2724_ =>
+ (sep_matches_prefix _s2724_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2725_ _) =>
+ (match (string_drop _s2724_ _s2725_) with
+ | _s2726_ =>
+ (frm_mnemonic_matches_prefix _s2726_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2727_ _) =>
+ let p0_ :=
+ string_drop _s2726_ _s2727_ in
+ if generic_eq p0_ "" then
+ Some (FSQRT_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2691_ (_s2692_ : string)
+: M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2692_ with
+ | _s2693_ =>
+ (f_bin_rm_type_mnemonic_D_matches_prefix _s2693_) >>= fun w__0 : option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s2694_ _) =>
+ (match (string_drop _s2693_ _s2694_) with
+ | _s2695_ =>
+ (spc_matches_prefix _s2695_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2696_ _) =>
+ (match (string_drop _s2695_ _s2696_) with
+ | _s2697_ =>
+ (freg_name_matches_prefix _s2697_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2698_ _) =>
+ (match (string_drop _s2697_ _s2698_) with
+ | _s2699_ =>
+ (sep_matches_prefix _s2699_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2700_ _) =>
+ (match (string_drop _s2699_ _s2700_) with
+ | _s2701_ =>
+ (freg_name_matches_prefix _s2701_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2702_ _) =>
+ (match (string_drop _s2701_ _s2702_) with
+ | _s2703_ =>
+ (sep_matches_prefix _s2703_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2704_ _) =>
+ (match (string_drop _s2703_ _s2704_) with
+ | _s2705_ =>
+ (freg_name_matches_prefix _s2705_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s2706_ _) =>
+ (match (string_drop _s2705_ _s2706_) with
+ | _s2707_ =>
+ (sep_matches_prefix _s2707_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s2708_ _) =>
+ (match (string_drop _s2707_
+ _s2708_) with
+ | _s2709_ =>
+ (frm_mnemonic_matches_prefix
+ _s2709_) >>= fun w__8 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__8 with
+ | Some
+ (rm, existT _ _s2710_ _) =>
+ let p0_ :=
+ string_drop
+ _s2709_
+ _s2710_ in
+ if generic_eq p0_
+ "" then
+ Some
+ (op, rd, rs1, rs2, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2666_ (_s2667_ : string)
+: M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2667_ with
+ | _s2668_ =>
+ (f_madd_type_mnemonic_D_matches_prefix _s2668_) >>= fun w__0 : option ((f_madd_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s2669_ _) =>
+ (match (string_drop _s2668_ _s2669_) with
+ | _s2670_ =>
+ (spc_matches_prefix _s2670_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2671_ _) =>
+ (match (string_drop _s2670_ _s2671_) with
+ | _s2672_ =>
+ (freg_name_matches_prefix _s2672_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2673_ _) =>
+ (match (string_drop _s2672_ _s2673_) with
+ | _s2674_ =>
+ (sep_matches_prefix _s2674_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2675_ _) =>
+ (match (string_drop _s2674_ _s2675_) with
+ | _s2676_ =>
+ (freg_name_matches_prefix _s2676_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2677_ _) =>
+ (match (string_drop _s2676_ _s2677_) with
+ | _s2678_ =>
+ (sep_matches_prefix _s2678_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2679_ _) =>
+ (match (string_drop _s2678_ _s2679_) with
+ | _s2680_ =>
+ (freg_name_matches_prefix _s2680_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s2681_ _) =>
+ (match (string_drop _s2680_ _s2681_) with
+ | _s2682_ =>
+ (sep_matches_prefix _s2682_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s2683_ _) =>
+ (match (string_drop _s2682_
+ _s2683_) with
+ | _s2684_ =>
+ (freg_name_matches_prefix
+ _s2684_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__8 with
+ | Some
+ (rs3, existT _ _s2685_ _) =>
+ (match (string_drop
+ _s2684_ _s2685_) with
+ | _s2686_ =>
+ (sep_matches_prefix
+ _s2686_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__9 with
+ | Some
+ (tt, existT _ _s2687_ _) =>
+ (match (string_drop
+ _s2686_
+ _s2687_) with
+ | _s2688_ =>
+ (frm_mnemonic_matches_prefix
+ _s2688_) >>= fun w__10 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__10 with
+ | Some
+ (rm, existT _ _s2689_ _) =>
+ let p0_ :=
+ string_drop
+ _s2688_
+ _s2689_ in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (op, rd, rs1, rs2, rs3, rm)
+ else
+ None
+ | _ =>
+ None
+ end)
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2650_ (_s2651_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s2652_ := _s2651_ in
+ (if string_startswith _s2652_ "c.fsw" then
+ (match (string_drop _s2652_ (projT1 (string_length "c.fsw"))) with
+ | _s2653_ =>
+ (spc_matches_prefix _s2653_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s2654_ _) =>
+ (match (string_drop _s2653_ _s2654_) with
+ | _s2655_ =>
+ (creg_name_matches_prefix _s2655_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rsc1, existT _ _s2656_ _) =>
+ (match (string_drop _s2655_ _s2656_) with
+ | _s2657_ =>
+ (sep_matches_prefix _s2657_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s2658_ _) =>
+ (match (string_drop _s2657_ _s2658_) with
+ | _s2659_ =>
+ (creg_name_matches_prefix _s2659_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc2, existT _ _s2660_ _) =>
+ (match (string_drop _s2659_ _s2660_) with
+ | _s2661_ =>
+ (sep_matches_prefix _s2661_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s2662_ _) =>
+ match (string_drop _s2661_ _s2662_) with
+ | _s2663_ =>
+ match (hex_bits_7_matches_prefix _s2663_) with
+ | Some (v__1342, existT _ _s2664_ _) =>
+ if eq_vec (subrange_vec_dec v__1342 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1342 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1342 6 2 in
+ let p0_ := string_drop _s2663_ _s2664_ in
+ if generic_eq p0_ "" then
+ Some (rsc1, rsc2, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5))).
+
+Definition _s2634_ (_s2635_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s2636_ := _s2635_ in
+ (if string_startswith _s2636_ "c.flw" then
+ (match (string_drop _s2636_ (projT1 (string_length "c.flw"))) with
+ | _s2637_ =>
+ (spc_matches_prefix _s2637_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s2638_ _) =>
+ (match (string_drop _s2637_ _s2638_) with
+ | _s2639_ =>
+ (creg_name_matches_prefix _s2639_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rdc, existT _ _s2640_ _) =>
+ (match (string_drop _s2639_ _s2640_) with
+ | _s2641_ =>
+ (sep_matches_prefix _s2641_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s2642_ _) =>
+ (match (string_drop _s2641_ _s2642_) with
+ | _s2643_ =>
+ (creg_name_matches_prefix _s2643_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc, existT _ _s2644_ _) =>
+ (match (string_drop _s2643_ _s2644_) with
+ | _s2645_ =>
+ (sep_matches_prefix _s2645_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s2646_ _) =>
+ match (string_drop _s2645_ _s2646_) with
+ | _s2647_ =>
+ match (hex_bits_7_matches_prefix _s2647_) with
+ | Some (v__1344, existT _ _s2648_ _) =>
+ if eq_vec (subrange_vec_dec v__1344 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1344 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1344 6 2 in
+ let p0_ := string_drop _s2647_ _s2648_ in
+ if generic_eq p0_ "" then
+ Some (rdc, rsc, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5))).
+
+Definition _s2622_ (_s2623_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s2624_ := _s2623_ in
+ (if string_startswith _s2624_ "c.fswsp" then
+ (match (string_drop _s2624_ (projT1 (string_length "c.fswsp"))) with
+ | _s2625_ =>
+ (spc_matches_prefix _s2625_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s2626_ _) =>
+ (match (string_drop _s2625_ _s2626_) with
+ | _s2627_ =>
+ (reg_name_matches_prefix _s2627_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s2628_ _) =>
+ (match (string_drop _s2627_ _s2628_) with
+ | _s2629_ =>
+ (sep_matches_prefix _s2629_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s2630_ _) =>
+ match (string_drop _s2629_ _s2630_) with
+ | _s2631_ =>
+ match (hex_bits_6_matches_prefix _s2631_) with
+ | Some (uimm, existT _ _s2632_ _) =>
+ let p0_ := string_drop _s2631_ _s2632_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6))).
+
+Definition _s2610_ (_s2611_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s2612_ := _s2611_ in
+ (if string_startswith _s2612_ "c.flwsp" then
+ (match (string_drop _s2612_ (projT1 (string_length "c.flwsp"))) with
+ | _s2613_ =>
+ (spc_matches_prefix _s2613_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s2614_ _) =>
+ (match (string_drop _s2613_ _s2614_) with
+ | _s2615_ =>
+ (reg_name_matches_prefix _s2615_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s2616_ _) =>
+ (match (string_drop _s2615_ _s2616_) with
+ | _s2617_ =>
+ (sep_matches_prefix _s2617_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s2618_ _) =>
+ match (string_drop _s2617_ _s2618_) with
+ | _s2619_ =>
+ match (hex_bits_6_matches_prefix _s2619_) with
+ | Some (imm, existT _ _s2620_ _) =>
+ let p0_ := string_drop _s2619_ _s2620_ in
+ if generic_eq p0_ "" then Some (rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6))).
+
+Definition _s2597_ (_s2598_ : string) : M (option ((f_un_op_S * mword 5 * mword 5))) :=
+ (match _s2598_ with
+ | _s2599_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s2599_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCLASS_S, existT _ _s2600_ _) =>
+ (match (string_drop _s2599_ _s2600_) with
+ | _s2601_ =>
+ (spc_matches_prefix _s2601_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2602_ _) =>
+ (match (string_drop _s2601_ _s2602_) with
+ | _s2603_ =>
+ (reg_name_matches_prefix _s2603_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2604_ _) =>
+ (match (string_drop _s2603_ _s2604_) with
+ | _s2605_ =>
+ (sep_matches_prefix _s2605_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2606_ _) =>
+ (match (string_drop _s2605_ _s2606_) with
+ | _s2607_ =>
+ (freg_name_matches_prefix _s2607_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s2608_ _) =>
+ let p0_ := string_drop _s2607_ _s2608_ in
+ if generic_eq p0_ "" then
+ Some (FCLASS_S, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5))).
+
+Definition _s2584_ (_s2585_ : string) : M (option ((f_un_op_S * mword 5 * mword 5))) :=
+ (match _s2585_ with
+ | _s2586_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s2586_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_W_X, existT _ _s2587_ _) =>
+ (match (string_drop _s2586_ _s2587_) with
+ | _s2588_ =>
+ (spc_matches_prefix _s2588_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2589_ _) =>
+ (match (string_drop _s2588_ _s2589_) with
+ | _s2590_ =>
+ (freg_name_matches_prefix _s2590_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2591_ _) =>
+ (match (string_drop _s2590_ _s2591_) with
+ | _s2592_ =>
+ (sep_matches_prefix _s2592_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2593_ _) =>
+ (match (string_drop _s2592_ _s2593_) with
+ | _s2594_ =>
+ (reg_name_matches_prefix _s2594_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s2595_ _) =>
+ let p0_ := string_drop _s2594_ _s2595_ in
+ if generic_eq p0_ "" then Some (FMV_W_X, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5))).
+
+Definition _s2571_ (_s2572_ : string) : M (option ((f_un_op_S * mword 5 * mword 5))) :=
+ (match _s2572_ with
+ | _s2573_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s2573_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_X_W, existT _ _s2574_ _) =>
+ (match (string_drop _s2573_ _s2574_) with
+ | _s2575_ =>
+ (spc_matches_prefix _s2575_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2576_ _) =>
+ (match (string_drop _s2575_ _s2576_) with
+ | _s2577_ =>
+ (reg_name_matches_prefix _s2577_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2578_ _) =>
+ (match (string_drop _s2577_ _s2578_) with
+ | _s2579_ =>
+ (sep_matches_prefix _s2579_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2580_ _) =>
+ (match (string_drop _s2579_ _s2580_) with
+ | _s2581_ =>
+ (freg_name_matches_prefix _s2581_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s2582_ _) =>
+ let p0_ := string_drop _s2581_ _s2582_ in
+ if generic_eq p0_ "" then Some (FMV_X_W, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5))).
+
+Definition _s2554_ (_s2555_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s2555_ with
+ | _s2556_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s2556_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLE_S, existT _ _s2557_ _) =>
+ (match (string_drop _s2556_ _s2557_) with
+ | _s2558_ =>
+ (spc_matches_prefix _s2558_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2559_ _) =>
+ (match (string_drop _s2558_ _s2559_) with
+ | _s2560_ =>
+ (reg_name_matches_prefix _s2560_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2561_ _) =>
+ (match (string_drop _s2560_ _s2561_) with
+ | _s2562_ =>
+ (sep_matches_prefix _s2562_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2563_ _) =>
+ (match (string_drop _s2562_ _s2563_) with
+ | _s2564_ =>
+ (freg_name_matches_prefix _s2564_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2565_ _) =>
+ (match (string_drop _s2564_ _s2565_) with
+ | _s2566_ =>
+ (sep_matches_prefix _s2566_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2567_ _) =>
+ (match (string_drop _s2566_ _s2567_) with
+ | _s2568_ =>
+ (freg_name_matches_prefix _s2568_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2569_ _) =>
+ let p0_ :=
+ string_drop _s2568_ _s2569_ in
+ if generic_eq p0_ "" then
+ Some (FLE_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s2537_ (_s2538_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s2538_ with
+ | _s2539_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s2539_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLT_S, existT _ _s2540_ _) =>
+ (match (string_drop _s2539_ _s2540_) with
+ | _s2541_ =>
+ (spc_matches_prefix _s2541_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2542_ _) =>
+ (match (string_drop _s2541_ _s2542_) with
+ | _s2543_ =>
+ (reg_name_matches_prefix _s2543_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2544_ _) =>
+ (match (string_drop _s2543_ _s2544_) with
+ | _s2545_ =>
+ (sep_matches_prefix _s2545_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2546_ _) =>
+ (match (string_drop _s2545_ _s2546_) with
+ | _s2547_ =>
+ (freg_name_matches_prefix _s2547_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2548_ _) =>
+ (match (string_drop _s2547_ _s2548_) with
+ | _s2549_ =>
+ (sep_matches_prefix _s2549_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2550_ _) =>
+ (match (string_drop _s2549_ _s2550_) with
+ | _s2551_ =>
+ (freg_name_matches_prefix _s2551_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2552_ _) =>
+ let p0_ :=
+ string_drop _s2551_ _s2552_ in
+ if generic_eq p0_ "" then
+ Some (FLT_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s2520_ (_s2521_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s2521_ with
+ | _s2522_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s2522_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FEQ_S, existT _ _s2523_ _) =>
+ (match (string_drop _s2522_ _s2523_) with
+ | _s2524_ =>
+ (spc_matches_prefix _s2524_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2525_ _) =>
+ (match (string_drop _s2524_ _s2525_) with
+ | _s2526_ =>
+ (reg_name_matches_prefix _s2526_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2527_ _) =>
+ (match (string_drop _s2526_ _s2527_) with
+ | _s2528_ =>
+ (sep_matches_prefix _s2528_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2529_ _) =>
+ (match (string_drop _s2528_ _s2529_) with
+ | _s2530_ =>
+ (freg_name_matches_prefix _s2530_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2531_ _) =>
+ (match (string_drop _s2530_ _s2531_) with
+ | _s2532_ =>
+ (sep_matches_prefix _s2532_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2533_ _) =>
+ (match (string_drop _s2532_ _s2533_) with
+ | _s2534_ =>
+ (freg_name_matches_prefix _s2534_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2535_ _) =>
+ let p0_ :=
+ string_drop _s2534_ _s2535_ in
+ if generic_eq p0_ "" then
+ Some (FEQ_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s2503_ (_s2504_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s2504_ with
+ | _s2505_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s2505_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMAX_S, existT _ _s2506_ _) =>
+ (match (string_drop _s2505_ _s2506_) with
+ | _s2507_ =>
+ (spc_matches_prefix _s2507_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2508_ _) =>
+ (match (string_drop _s2507_ _s2508_) with
+ | _s2509_ =>
+ (freg_name_matches_prefix _s2509_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2510_ _) =>
+ (match (string_drop _s2509_ _s2510_) with
+ | _s2511_ =>
+ (sep_matches_prefix _s2511_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2512_ _) =>
+ (match (string_drop _s2511_ _s2512_) with
+ | _s2513_ =>
+ (freg_name_matches_prefix _s2513_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2514_ _) =>
+ (match (string_drop _s2513_ _s2514_) with
+ | _s2515_ =>
+ (sep_matches_prefix _s2515_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2516_ _) =>
+ (match (string_drop _s2515_ _s2516_) with
+ | _s2517_ =>
+ (freg_name_matches_prefix _s2517_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2518_ _) =>
+ let p0_ :=
+ string_drop _s2517_ _s2518_ in
+ if generic_eq p0_ "" then
+ Some (FMAX_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s2486_ (_s2487_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s2487_ with
+ | _s2488_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s2488_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMIN_S, existT _ _s2489_ _) =>
+ (match (string_drop _s2488_ _s2489_) with
+ | _s2490_ =>
+ (spc_matches_prefix _s2490_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2491_ _) =>
+ (match (string_drop _s2490_ _s2491_) with
+ | _s2492_ =>
+ (freg_name_matches_prefix _s2492_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2493_ _) =>
+ (match (string_drop _s2492_ _s2493_) with
+ | _s2494_ =>
+ (sep_matches_prefix _s2494_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2495_ _) =>
+ (match (string_drop _s2494_ _s2495_) with
+ | _s2496_ =>
+ (freg_name_matches_prefix _s2496_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2497_ _) =>
+ (match (string_drop _s2496_ _s2497_) with
+ | _s2498_ =>
+ (sep_matches_prefix _s2498_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2499_ _) =>
+ (match (string_drop _s2498_ _s2499_) with
+ | _s2500_ =>
+ (freg_name_matches_prefix _s2500_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2501_ _) =>
+ let p0_ :=
+ string_drop _s2500_ _s2501_ in
+ if generic_eq p0_ "" then
+ Some (FMIN_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s2469_ (_s2470_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s2470_ with
+ | _s2471_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s2471_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJX_S, existT _ _s2472_ _) =>
+ (match (string_drop _s2471_ _s2472_) with
+ | _s2473_ =>
+ (spc_matches_prefix _s2473_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2474_ _) =>
+ (match (string_drop _s2473_ _s2474_) with
+ | _s2475_ =>
+ (freg_name_matches_prefix _s2475_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2476_ _) =>
+ (match (string_drop _s2475_ _s2476_) with
+ | _s2477_ =>
+ (sep_matches_prefix _s2477_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2478_ _) =>
+ (match (string_drop _s2477_ _s2478_) with
+ | _s2479_ =>
+ (freg_name_matches_prefix _s2479_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2480_ _) =>
+ (match (string_drop _s2479_ _s2480_) with
+ | _s2481_ =>
+ (sep_matches_prefix _s2481_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2482_ _) =>
+ (match (string_drop _s2481_ _s2482_) with
+ | _s2483_ =>
+ (freg_name_matches_prefix _s2483_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2484_ _) =>
+ let p0_ :=
+ string_drop _s2483_ _s2484_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJX_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s2452_ (_s2453_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s2453_ with
+ | _s2454_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s2454_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJN_S, existT _ _s2455_ _) =>
+ (match (string_drop _s2454_ _s2455_) with
+ | _s2456_ =>
+ (spc_matches_prefix _s2456_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2457_ _) =>
+ (match (string_drop _s2456_ _s2457_) with
+ | _s2458_ =>
+ (freg_name_matches_prefix _s2458_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2459_ _) =>
+ (match (string_drop _s2458_ _s2459_) with
+ | _s2460_ =>
+ (sep_matches_prefix _s2460_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2461_ _) =>
+ (match (string_drop _s2460_ _s2461_) with
+ | _s2462_ =>
+ (freg_name_matches_prefix _s2462_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2463_ _) =>
+ (match (string_drop _s2462_ _s2463_) with
+ | _s2464_ =>
+ (sep_matches_prefix _s2464_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2465_ _) =>
+ (match (string_drop _s2464_ _s2465_) with
+ | _s2466_ =>
+ (freg_name_matches_prefix _s2466_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2467_ _) =>
+ let p0_ :=
+ string_drop _s2466_ _s2467_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJN_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s2435_ (_s2436_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s2436_ with
+ | _s2437_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s2437_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJ_S, existT _ _s2438_ _) =>
+ (match (string_drop _s2437_ _s2438_) with
+ | _s2439_ =>
+ (spc_matches_prefix _s2439_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2440_ _) =>
+ (match (string_drop _s2439_ _s2440_) with
+ | _s2441_ =>
+ (freg_name_matches_prefix _s2441_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2442_ _) =>
+ (match (string_drop _s2441_ _s2442_) with
+ | _s2443_ =>
+ (sep_matches_prefix _s2443_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2444_ _) =>
+ (match (string_drop _s2443_ _s2444_) with
+ | _s2445_ =>
+ (freg_name_matches_prefix _s2445_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2446_ _) =>
+ (match (string_drop _s2445_ _s2446_) with
+ | _s2447_ =>
+ (sep_matches_prefix _s2447_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2448_ _) =>
+ (match (string_drop _s2447_ _s2448_) with
+ | _s2449_ =>
+ (freg_name_matches_prefix _s2449_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2450_ _) =>
+ let p0_ :=
+ string_drop _s2449_ _s2450_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJ_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s2418_ (_s2419_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2419_ with
+ | _s2420_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2420_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_LU, existT _ _s2421_ _) =>
+ (match (string_drop _s2420_ _s2421_) with
+ | _s2422_ =>
+ (spc_matches_prefix _s2422_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2423_ _) =>
+ (match (string_drop _s2422_ _s2423_) with
+ | _s2424_ =>
+ (freg_name_matches_prefix _s2424_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2425_ _) =>
+ (match (string_drop _s2424_ _s2425_) with
+ | _s2426_ =>
+ (sep_matches_prefix _s2426_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2427_ _) =>
+ (match (string_drop _s2426_ _s2427_) with
+ | _s2428_ =>
+ (reg_name_matches_prefix _s2428_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2429_ _) =>
+ (match (string_drop _s2428_ _s2429_) with
+ | _s2430_ =>
+ (sep_matches_prefix _s2430_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2431_ _) =>
+ (match (string_drop _s2430_ _s2431_) with
+ | _s2432_ =>
+ (frm_mnemonic_matches_prefix _s2432_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2433_ _) =>
+ let p0_ :=
+ string_drop _s2432_ _s2433_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_LU, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2401_ (_s2402_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2402_ with
+ | _s2403_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2403_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_L, existT _ _s2404_ _) =>
+ (match (string_drop _s2403_ _s2404_) with
+ | _s2405_ =>
+ (spc_matches_prefix _s2405_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2406_ _) =>
+ (match (string_drop _s2405_ _s2406_) with
+ | _s2407_ =>
+ (freg_name_matches_prefix _s2407_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2408_ _) =>
+ (match (string_drop _s2407_ _s2408_) with
+ | _s2409_ =>
+ (sep_matches_prefix _s2409_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2410_ _) =>
+ (match (string_drop _s2409_ _s2410_) with
+ | _s2411_ =>
+ (reg_name_matches_prefix _s2411_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2412_ _) =>
+ (match (string_drop _s2411_ _s2412_) with
+ | _s2413_ =>
+ (sep_matches_prefix _s2413_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2414_ _) =>
+ (match (string_drop _s2413_ _s2414_) with
+ | _s2415_ =>
+ (frm_mnemonic_matches_prefix _s2415_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2416_ _) =>
+ let p0_ :=
+ string_drop _s2415_ _s2416_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_L, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2384_ (_s2385_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2385_ with
+ | _s2386_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2386_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_LU_S, existT _ _s2387_ _) =>
+ (match (string_drop _s2386_ _s2387_) with
+ | _s2388_ =>
+ (spc_matches_prefix _s2388_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2389_ _) =>
+ (match (string_drop _s2388_ _s2389_) with
+ | _s2390_ =>
+ (reg_name_matches_prefix _s2390_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2391_ _) =>
+ (match (string_drop _s2390_ _s2391_) with
+ | _s2392_ =>
+ (sep_matches_prefix _s2392_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2393_ _) =>
+ (match (string_drop _s2392_ _s2393_) with
+ | _s2394_ =>
+ (freg_name_matches_prefix _s2394_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2395_ _) =>
+ (match (string_drop _s2394_ _s2395_) with
+ | _s2396_ =>
+ (sep_matches_prefix _s2396_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2397_ _) =>
+ (match (string_drop _s2396_ _s2397_) with
+ | _s2398_ =>
+ (frm_mnemonic_matches_prefix _s2398_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2399_ _) =>
+ let p0_ :=
+ string_drop _s2398_ _s2399_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_LU_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2367_ (_s2368_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2368_ with
+ | _s2369_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2369_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_L_S, existT _ _s2370_ _) =>
+ (match (string_drop _s2369_ _s2370_) with
+ | _s2371_ =>
+ (spc_matches_prefix _s2371_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2372_ _) =>
+ (match (string_drop _s2371_ _s2372_) with
+ | _s2373_ =>
+ (reg_name_matches_prefix _s2373_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2374_ _) =>
+ (match (string_drop _s2373_ _s2374_) with
+ | _s2375_ =>
+ (sep_matches_prefix _s2375_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2376_ _) =>
+ (match (string_drop _s2375_ _s2376_) with
+ | _s2377_ =>
+ (freg_name_matches_prefix _s2377_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2378_ _) =>
+ (match (string_drop _s2377_ _s2378_) with
+ | _s2379_ =>
+ (sep_matches_prefix _s2379_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2380_ _) =>
+ (match (string_drop _s2379_ _s2380_) with
+ | _s2381_ =>
+ (frm_mnemonic_matches_prefix _s2381_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2382_ _) =>
+ let p0_ :=
+ string_drop _s2381_ _s2382_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_L_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2350_ (_s2351_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2351_ with
+ | _s2352_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2352_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_WU, existT _ _s2353_ _) =>
+ (match (string_drop _s2352_ _s2353_) with
+ | _s2354_ =>
+ (spc_matches_prefix _s2354_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2355_ _) =>
+ (match (string_drop _s2354_ _s2355_) with
+ | _s2356_ =>
+ (freg_name_matches_prefix _s2356_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2357_ _) =>
+ (match (string_drop _s2356_ _s2357_) with
+ | _s2358_ =>
+ (sep_matches_prefix _s2358_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2359_ _) =>
+ (match (string_drop _s2358_ _s2359_) with
+ | _s2360_ =>
+ (reg_name_matches_prefix _s2360_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2361_ _) =>
+ (match (string_drop _s2360_ _s2361_) with
+ | _s2362_ =>
+ (sep_matches_prefix _s2362_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2363_ _) =>
+ (match (string_drop _s2362_ _s2363_) with
+ | _s2364_ =>
+ (frm_mnemonic_matches_prefix _s2364_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2365_ _) =>
+ let p0_ :=
+ string_drop _s2364_ _s2365_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_WU, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2333_ (_s2334_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2334_ with
+ | _s2335_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2335_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_W, existT _ _s2336_ _) =>
+ (match (string_drop _s2335_ _s2336_) with
+ | _s2337_ =>
+ (spc_matches_prefix _s2337_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2338_ _) =>
+ (match (string_drop _s2337_ _s2338_) with
+ | _s2339_ =>
+ (freg_name_matches_prefix _s2339_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2340_ _) =>
+ (match (string_drop _s2339_ _s2340_) with
+ | _s2341_ =>
+ (sep_matches_prefix _s2341_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2342_ _) =>
+ (match (string_drop _s2341_ _s2342_) with
+ | _s2343_ =>
+ (reg_name_matches_prefix _s2343_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2344_ _) =>
+ (match (string_drop _s2343_ _s2344_) with
+ | _s2345_ =>
+ (sep_matches_prefix _s2345_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2346_ _) =>
+ (match (string_drop _s2345_ _s2346_) with
+ | _s2347_ =>
+ (frm_mnemonic_matches_prefix _s2347_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2348_ _) =>
+ let p0_ :=
+ string_drop _s2347_ _s2348_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_W, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2316_ (_s2317_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2317_ with
+ | _s2318_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2318_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_WU_S, existT _ _s2319_ _) =>
+ (match (string_drop _s2318_ _s2319_) with
+ | _s2320_ =>
+ (spc_matches_prefix _s2320_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2321_ _) =>
+ (match (string_drop _s2320_ _s2321_) with
+ | _s2322_ =>
+ (reg_name_matches_prefix _s2322_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2323_ _) =>
+ (match (string_drop _s2322_ _s2323_) with
+ | _s2324_ =>
+ (sep_matches_prefix _s2324_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2325_ _) =>
+ (match (string_drop _s2324_ _s2325_) with
+ | _s2326_ =>
+ (freg_name_matches_prefix _s2326_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2327_ _) =>
+ (match (string_drop _s2326_ _s2327_) with
+ | _s2328_ =>
+ (sep_matches_prefix _s2328_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2329_ _) =>
+ (match (string_drop _s2328_ _s2329_) with
+ | _s2330_ =>
+ (frm_mnemonic_matches_prefix _s2330_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2331_ _) =>
+ let p0_ :=
+ string_drop _s2330_ _s2331_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_WU_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2299_ (_s2300_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2300_ with
+ | _s2301_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2301_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_W_S, existT _ _s2302_ _) =>
+ (match (string_drop _s2301_ _s2302_) with
+ | _s2303_ =>
+ (spc_matches_prefix _s2303_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2304_ _) =>
+ (match (string_drop _s2303_ _s2304_) with
+ | _s2305_ =>
+ (reg_name_matches_prefix _s2305_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2306_ _) =>
+ (match (string_drop _s2305_ _s2306_) with
+ | _s2307_ =>
+ (sep_matches_prefix _s2307_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2308_ _) =>
+ (match (string_drop _s2307_ _s2308_) with
+ | _s2309_ =>
+ (freg_name_matches_prefix _s2309_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2310_ _) =>
+ (match (string_drop _s2309_ _s2310_) with
+ | _s2311_ =>
+ (sep_matches_prefix _s2311_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2312_ _) =>
+ (match (string_drop _s2311_ _s2312_) with
+ | _s2313_ =>
+ (frm_mnemonic_matches_prefix _s2313_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2314_ _) =>
+ let p0_ :=
+ string_drop _s2313_ _s2314_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_W_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2282_ (_s2283_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2283_ with
+ | _s2284_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s2284_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSQRT_S, existT _ _s2285_ _) =>
+ (match (string_drop _s2284_ _s2285_) with
+ | _s2286_ =>
+ (spc_matches_prefix _s2286_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2287_ _) =>
+ (match (string_drop _s2286_ _s2287_) with
+ | _s2288_ =>
+ (freg_name_matches_prefix _s2288_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2289_ _) =>
+ (match (string_drop _s2288_ _s2289_) with
+ | _s2290_ =>
+ (sep_matches_prefix _s2290_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2291_ _) =>
+ (match (string_drop _s2290_ _s2291_) with
+ | _s2292_ =>
+ (freg_name_matches_prefix _s2292_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2293_ _) =>
+ (match (string_drop _s2292_ _s2293_) with
+ | _s2294_ =>
+ (sep_matches_prefix _s2294_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2295_ _) =>
+ (match (string_drop _s2294_ _s2295_) with
+ | _s2296_ =>
+ (frm_mnemonic_matches_prefix _s2296_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s2297_ _) =>
+ let p0_ :=
+ string_drop _s2296_ _s2297_ in
+ if generic_eq p0_ "" then
+ Some (FSQRT_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2261_ (_s2262_ : string)
+: M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2262_ with
+ | _s2263_ =>
+ (f_bin_rm_type_mnemonic_S_matches_prefix _s2263_) >>= fun w__0 : option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s2264_ _) =>
+ (match (string_drop _s2263_ _s2264_) with
+ | _s2265_ =>
+ (spc_matches_prefix _s2265_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2266_ _) =>
+ (match (string_drop _s2265_ _s2266_) with
+ | _s2267_ =>
+ (freg_name_matches_prefix _s2267_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2268_ _) =>
+ (match (string_drop _s2267_ _s2268_) with
+ | _s2269_ =>
+ (sep_matches_prefix _s2269_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2270_ _) =>
+ (match (string_drop _s2269_ _s2270_) with
+ | _s2271_ =>
+ (freg_name_matches_prefix _s2271_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2272_ _) =>
+ (match (string_drop _s2271_ _s2272_) with
+ | _s2273_ =>
+ (sep_matches_prefix _s2273_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2274_ _) =>
+ (match (string_drop _s2273_ _s2274_) with
+ | _s2275_ =>
+ (freg_name_matches_prefix _s2275_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s2276_ _) =>
+ (match (string_drop _s2275_ _s2276_) with
+ | _s2277_ =>
+ (sep_matches_prefix _s2277_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s2278_ _) =>
+ (match (string_drop _s2277_
+ _s2278_) with
+ | _s2279_ =>
+ (frm_mnemonic_matches_prefix
+ _s2279_) >>= fun w__8 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__8 with
+ | Some
+ (rm, existT _ _s2280_ _) =>
+ let p0_ :=
+ string_drop
+ _s2279_
+ _s2280_ in
+ if generic_eq p0_
+ "" then
+ Some
+ (op, rd, rs1, rs2, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12))).
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode))).
-Definition _s1634_ (_s1635_ : string)
-: M (option ((csrop * mword 5 * mword 5 * mword 12))) :=
-
- (match _s1635_ with
- | _s1636_ =>
- (csr_mnemonic_matches_prefix _s1636_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=
+Definition _s2236_ (_s2237_ : string)
+: M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s2237_ with
+ | _s2238_ =>
+ (f_madd_type_mnemonic_S_matches_prefix _s2238_) >>= fun w__0 : option ((f_madd_op_S * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1637_ _)) =>
- let _s1638_ := string_drop _s1636_ _s1637_ in
- (if ((string_startswith _s1638_ "i")) then
- (match (string_drop _s1638_ (projT1 (string_length "i"))) with
- | _s1639_ =>
- (spc_matches_prefix _s1639_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s2239_ _) =>
+ (match (string_drop _s2238_ _s2239_) with
+ | _s2240_ =>
+ (spc_matches_prefix _s2240_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2241_ _) =>
+ (match (string_drop _s2240_ _s2241_) with
+ | _s2242_ =>
+ (freg_name_matches_prefix _s2242_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2243_ _) =>
+ (match (string_drop _s2242_ _s2243_) with
+ | _s2244_ =>
+ (sep_matches_prefix _s2244_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2245_ _) =>
+ (match (string_drop _s2244_ _s2245_) with
+ | _s2246_ =>
+ (freg_name_matches_prefix _s2246_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s2247_ _) =>
+ (match (string_drop _s2246_ _s2247_) with
+ | _s2248_ =>
+ (sep_matches_prefix _s2248_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2249_ _) =>
+ (match (string_drop _s2248_ _s2249_) with
+ | _s2250_ =>
+ (freg_name_matches_prefix _s2250_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s2251_ _) =>
+ (match (string_drop _s2250_ _s2251_) with
+ | _s2252_ =>
+ (sep_matches_prefix _s2252_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s2253_ _) =>
+ (match (string_drop _s2252_
+ _s2253_) with
+ | _s2254_ =>
+ (freg_name_matches_prefix
+ _s2254_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__8 with
+ | Some
+ (rs3, existT _ _s2255_ _) =>
+ (match (string_drop
+ _s2254_ _s2255_) with
+ | _s2256_ =>
+ (sep_matches_prefix
+ _s2256_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__9 with
+ | Some
+ (tt, existT _ _s2257_ _) =>
+ (match (string_drop
+ _s2256_
+ _s2257_) with
+ | _s2258_ =>
+ (frm_mnemonic_matches_prefix
+ _s2258_) >>= fun w__10 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__10 with
+ | Some
+ (rm, existT _ _s2259_ _) =>
+ let p0_ :=
+ string_drop
+ _s2258_
+ _s2259_ in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (op, rd, rs1, rs2, rs3, rm)
+ else
+ None
+ | _ =>
+ None
+ end)
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s2212_ (_s2213_ : string) : M (option ((word_width * mword 5 * mword 12 * mword 5))) :=
+ let _s2214_ := _s2213_ in
+ (if string_startswith _s2214_ "fs" then
+ (match (string_drop _s2214_ (projT1 (string_length "fs"))) with
+ | _s2215_ =>
+ (size_mnemonic_matches_prefix _s2215_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (width, existT _ _s2216_ _) =>
+ (match (string_drop _s2215_ _s2216_) with
+ | _s2217_ =>
+ (spc_matches_prefix _s2217_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1640_ _)) =>
- (match (string_drop _s1639_ _s1640_) with
- | _s1641_ =>
- (reg_name_matches_prefix _s1641_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2218_ _) =>
+ (match (string_drop _s2217_ _s2218_) with
+ | _s2219_ =>
+ (freg_name_matches_prefix _s2219_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1642_ _)) =>
- (match (string_drop _s1641_ _s1642_) with
- | _s1643_ =>
- (sep_matches_prefix _s1643_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s2220_ _) =>
+ (match (string_drop _s2219_ _s2220_) with
+ | _s2221_ =>
+ (sep_matches_prefix _s2221_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1644_ _)) =>
- (match (string_drop _s1643_ _s1644_) with
- | _s1645_ =>
- (match (hex_bits_5_matches_prefix _s1645_) with
- | Some ((rs1, existT _ _s1646_ _)) =>
- (match (string_drop _s1645_ _s1646_) with
- | _s1647_ =>
- (sep_matches_prefix _s1647_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2222_ _) =>
+ (match (string_drop _s2221_ _s2222_) with
+ | _s2223_ =>
+ (match (hex_bits_12_matches_prefix _s2223_) with
+ | Some (imm, existT _ _s2224_ _) =>
+ (match (string_drop _s2223_ _s2224_) with
+ | _s2225_ =>
+ (opt_spc_matches_prefix _s2225_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s1648_ _)) =>
- (match (string_drop _s1647_ _s1648_) with
- | _s1649_ =>
- (csr_name_map_matches_prefix _s1649_) >>= fun w__5 : option ((mword 12 * {n : Z & ArithFact (n >=
- 0)})) =>
- returnm ((match w__5 with
- | Some
- ((csr, existT _ _s1650_ _)) =>
- let p0_ :=
- string_drop _s1649_ _s1650_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, csr))
- else None
- | _ => None
- end)
- : option ((csrop * mword 5 * mword 5 * mword 12)))
- end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ | Some (tt, existT _ _s2226_ _) =>
+ let _s2227_ := string_drop _s2225_ _s2226_ in
+ (if string_startswith _s2227_ "(" then
+ (match (string_drop _s2227_
+ (projT1
+ (string_length "("))) with
+ | _s2228_ =>
+ (opt_spc_matches_prefix _s2228_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2229_ _) =>
+ (match (string_drop _s2228_ _s2229_) with
+ | _s2230_ =>
+ (reg_name_matches_prefix
+ _s2230_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some
+ (rs1, existT _ _s2231_ _) =>
+ (match (string_drop _s2230_
+ _s2231_) with
+ | _s2232_ =>
+ (opt_spc_matches_prefix
+ _s2232_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__7 with
+ | Some
+ (tt, existT _ _s2233_ _) =>
+ let _s2234_ :=
+ string_drop
+ _s2232_
+ _s2233_ in
+ if string_startswith
+ _s2234_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s2234_
+ (projT1
+ (string_length
+ ")")) in
+ if generic_eq
+ p0_
+ ""
+ then
+ Some
+ (width, rs2, imm, rs1)
+ else None
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5))).
+
+Definition _s2188_ (_s2189_ : string) : M (option ((word_width * mword 5 * mword 12 * mword 5))) :=
+ let _s2190_ := _s2189_ in
+ (if string_startswith _s2190_ "fl" then
+ (match (string_drop _s2190_ (projT1 (string_length "fl"))) with
+ | _s2191_ =>
+ (size_mnemonic_matches_prefix _s2191_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (width, existT _ _s2192_ _) =>
+ (match (string_drop _s2191_ _s2192_) with
+ | _s2193_ =>
+ (spc_matches_prefix _s2193_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2194_ _) =>
+ (match (string_drop _s2193_ _s2194_) with
+ | _s2195_ =>
+ (freg_name_matches_prefix _s2195_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2196_ _) =>
+ (match (string_drop _s2195_ _s2196_) with
+ | _s2197_ =>
+ (sep_matches_prefix _s2197_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2198_ _) =>
+ (match (string_drop _s2197_ _s2198_) with
+ | _s2199_ =>
+ (match (hex_bits_12_matches_prefix _s2199_) with
+ | Some (imm, existT _ _s2200_ _) =>
+ (match (string_drop _s2199_ _s2200_) with
+ | _s2201_ =>
+ (opt_spc_matches_prefix _s2201_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (tt, existT _ _s2202_ _) =>
+ let _s2203_ := string_drop _s2201_ _s2202_ in
+ (if string_startswith _s2203_ "(" then
+ (match (string_drop _s2203_
+ (projT1
+ (string_length "("))) with
+ | _s2204_ =>
+ (opt_spc_matches_prefix _s2204_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2205_ _) =>
+ (match (string_drop _s2204_ _s2205_) with
+ | _s2206_ =>
+ (reg_name_matches_prefix
+ _s2206_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some
+ (rs1, existT _ _s2207_ _) =>
+ (match (string_drop _s2206_
+ _s2207_) with
+ | _s2208_ =>
+ (opt_spc_matches_prefix
+ _s2208_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__7 with
+ | Some
+ (tt, existT _ _s2209_ _) =>
+ let _s2210_ :=
+ string_drop
+ _s2208_
+ _s2209_ in
+ if string_startswith
+ _s2210_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s2210_
+ (projT1
+ (string_length
+ ")")) in
+ if generic_eq
+ p0_
+ ""
+ then
+ Some
+ (width, rd, imm, rs1)
+ else None
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5))).
+
+Definition _s2171_ (_s2172_ : string) : M (option ((csrop * mword 5 * mword 12 * mword 5))) :=
+ (match _s2172_ with
+ | _s2173_ =>
+ (csr_mnemonic_matches_prefix _s2173_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s2174_ _) =>
+ (match (string_drop _s2173_ _s2174_) with
+ | _s2175_ =>
+ (spc_matches_prefix _s2175_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2176_ _) =>
+ (match (string_drop _s2175_ _s2176_) with
+ | _s2177_ =>
+ (reg_name_matches_prefix _s2177_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2178_ _) =>
+ (match (string_drop _s2177_ _s2178_) with
+ | _s2179_ =>
+ (sep_matches_prefix _s2179_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2180_ _) =>
+ (match (string_drop _s2179_ _s2180_) with
+ | _s2181_ =>
+ (csr_name_map_matches_prefix _s2181_) >>= fun w__4 : option ((mword 12 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (csr, existT _ _s2182_ _) =>
+ (match (string_drop _s2181_ _s2182_) with
+ | _s2183_ =>
+ (sep_matches_prefix _s2183_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s2184_ _) =>
+ (match (string_drop _s2183_ _s2184_) with
+ | _s2185_ =>
+ (reg_name_matches_prefix _s2185_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs1, existT _ _s2186_ _) =>
+ let p0_ :=
+ string_drop _s2185_ _s2186_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, csr, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5))).
+
+Definition _s2153_ (_s2154_ : string) : M (option ((csrop * mword 5 * mword 12 * mword 5))) :=
+ (match _s2154_ with
+ | _s2155_ =>
+ (csr_mnemonic_matches_prefix _s2155_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s2156_ _) =>
+ let _s2157_ := string_drop _s2155_ _s2156_ in
+ (if string_startswith _s2157_ "i" then
+ (match (string_drop _s2157_ (projT1 (string_length "i"))) with
+ | _s2158_ =>
+ (spc_matches_prefix _s2158_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s2159_ _) =>
+ (match (string_drop _s2158_ _s2159_) with
+ | _s2160_ =>
+ (reg_name_matches_prefix _s2160_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s2161_ _) =>
+ (match (string_drop _s2160_ _s2161_) with
+ | _s2162_ =>
+ (sep_matches_prefix _s2162_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s2163_ _) =>
+ (match (string_drop _s2162_ _s2163_) with
+ | _s2164_ =>
+ (csr_name_map_matches_prefix _s2164_) >>= fun w__4 : option ((mword 12 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (csr, existT _ _s2165_ _) =>
+ (match (string_drop _s2164_ _s2165_) with
+ | _s2166_ =>
+ (sep_matches_prefix _s2166_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__5 with
+ | Some (tt, existT _ _s2167_ _) =>
+ match (string_drop _s2166_ _s2167_) with
+ | _s2168_ =>
+ match (hex_bits_5_matches_prefix
+ _s2168_) with
+ | Some (rs1, existT _ _s2169_ _) =>
+ let p0_ :=
+ string_drop _s2168_ _s2169_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, csr, rs1)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- else returnm (None : option ((csrop * mword 5 * mword 5 * mword 12))))
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12))).
-
-Definition _s1615_ (_s1616_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s1617_ := _s1616_ in
- (if ((string_startswith _s1617_ "rem")) then
- (match (string_drop _s1617_ (projT1 (string_length "rem"))) with
- | _s1618_ =>
- (maybe_not_u_matches_prefix _s1618_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5))).
+
+Definition _s2134_ (_s2135_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5))) :=
+ let _s2136_ := _s2135_ in
+ (if string_startswith _s2136_ "rem" then
+ (match (string_drop _s2136_ (projT1 (string_length "rem"))) with
+ | _s2137_ =>
+ (maybe_not_u_matches_prefix _s2137_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s1619_ _)) =>
- let _s1620_ := string_drop _s1618_ _s1619_ in
- (if ((string_startswith _s1620_ "w")) then
- (match (string_drop _s1620_ (projT1 (string_length "w"))) with
- | _s1621_ =>
- (spc_matches_prefix _s1621_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s2138_ _) =>
+ let _s2139_ := string_drop _s2137_ _s2138_ in
+ (if string_startswith _s2139_ "w" then
+ (match (string_drop _s2139_ (projT1 (string_length "w"))) with
+ | _s2140_ =>
+ (spc_matches_prefix _s2140_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1622_ _)) =>
- (match (string_drop _s1621_ _s1622_) with
- | _s1623_ =>
- (reg_name_matches_prefix _s1623_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2141_ _) =>
+ (match (string_drop _s2140_ _s2141_) with
+ | _s2142_ =>
+ (reg_name_matches_prefix _s2142_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1624_ _)) =>
- (match (string_drop _s1623_ _s1624_) with
- | _s1625_ =>
- (sep_matches_prefix _s1625_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s2143_ _) =>
+ (match (string_drop _s2142_ _s2143_) with
+ | _s2144_ =>
+ (sep_matches_prefix _s2144_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1626_ _)) =>
- (match (string_drop _s1625_ _s1626_) with
- | _s1627_ =>
- (reg_name_matches_prefix _s1627_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2145_ _) =>
+ (match (string_drop _s2144_ _s2145_) with
+ | _s2146_ =>
+ (reg_name_matches_prefix _s2146_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1628_ _)) =>
- (match (string_drop _s1627_ _s1628_) with
- | _s1629_ =>
- (sep_matches_prefix _s1629_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s2147_ _) =>
+ (match (string_drop _s2146_ _s2147_) with
+ | _s2148_ =>
+ (sep_matches_prefix _s2148_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1630_ _)) =>
- (match (string_drop _s1629_ _s1630_) with
- | _s1631_ =>
- (reg_name_matches_prefix _s1631_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2149_ _) =>
+ (match (string_drop _s2148_ _s2149_) with
+ | _s2150_ =>
+ (reg_name_matches_prefix _s2150_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s1632_ _)) =>
- let p0_ :=
- string_drop _s1631_
- _s1632_ in
- if ((generic_eq p0_ ""))
- then
- Some
- ((s, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some
+ (rs2, existT _ _s2151_ _) =>
+ let p0_ :=
+ string_drop _s2150_
+ _s2151_ in
+ if generic_eq p0_ "" then
+ Some (s, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5))).
-Definition _s1596_ (_s1597_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s1598_ := _s1597_ in
- (if ((string_startswith _s1598_ "div")) then
- (match (string_drop _s1598_ (projT1 (string_length "div"))) with
- | _s1599_ =>
- (maybe_not_u_matches_prefix _s1599_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s2115_ (_s2116_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5))) :=
+ let _s2117_ := _s2116_ in
+ (if string_startswith _s2117_ "div" then
+ (match (string_drop _s2117_ (projT1 (string_length "div"))) with
+ | _s2118_ =>
+ (maybe_not_u_matches_prefix _s2118_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s1600_ _)) =>
- let _s1601_ := string_drop _s1599_ _s1600_ in
- (if ((string_startswith _s1601_ "w")) then
- (match (string_drop _s1601_ (projT1 (string_length "w"))) with
- | _s1602_ =>
- (spc_matches_prefix _s1602_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s2119_ _) =>
+ let _s2120_ := string_drop _s2118_ _s2119_ in
+ (if string_startswith _s2120_ "w" then
+ (match (string_drop _s2120_ (projT1 (string_length "w"))) with
+ | _s2121_ =>
+ (spc_matches_prefix _s2121_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1603_ _)) =>
- (match (string_drop _s1602_ _s1603_) with
- | _s1604_ =>
- (reg_name_matches_prefix _s1604_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2122_ _) =>
+ (match (string_drop _s2121_ _s2122_) with
+ | _s2123_ =>
+ (reg_name_matches_prefix _s2123_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1605_ _)) =>
- (match (string_drop _s1604_ _s1605_) with
- | _s1606_ =>
- (sep_matches_prefix _s1606_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s2124_ _) =>
+ (match (string_drop _s2123_ _s2124_) with
+ | _s2125_ =>
+ (sep_matches_prefix _s2125_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1607_ _)) =>
- (match (string_drop _s1606_ _s1607_) with
- | _s1608_ =>
- (reg_name_matches_prefix _s1608_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2126_ _) =>
+ (match (string_drop _s2125_ _s2126_) with
+ | _s2127_ =>
+ (reg_name_matches_prefix _s2127_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1609_ _)) =>
- (match (string_drop _s1608_ _s1609_) with
- | _s1610_ =>
- (sep_matches_prefix _s1610_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s2128_ _) =>
+ (match (string_drop _s2127_ _s2128_) with
+ | _s2129_ =>
+ (sep_matches_prefix _s2129_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1611_ _)) =>
- (match (string_drop _s1610_ _s1611_) with
- | _s1612_ =>
- (reg_name_matches_prefix _s1612_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2130_ _) =>
+ (match (string_drop _s2129_ _s2130_) with
+ | _s2131_ =>
+ (reg_name_matches_prefix _s2131_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s1613_ _)) =>
- let p0_ :=
- string_drop _s1612_
- _s1613_ in
- if ((generic_eq p0_ ""))
- then
- Some
- ((s, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some
+ (rs2, existT _ _s2132_ _) =>
+ let p0_ :=
+ string_drop _s2131_
+ _s2132_ in
+ if generic_eq p0_ "" then
+ Some (s, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5))).
-Definition _s1580_ (_s1581_ : string)
-: M (option ((mword 5 * mword 5 * mword 5))) :=
-
- let _s1582_ := _s1581_ in
- (if ((string_startswith _s1582_ "mulw")) then
- (match (string_drop _s1582_ (projT1 (string_length "mulw"))) with
- | _s1583_ =>
- (spc_matches_prefix _s1583_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s2099_ (_s2100_ : string) : M (option ((mword 5 * mword 5 * mword 5))) :=
+ let _s2101_ := _s2100_ in
+ (if string_startswith _s2101_ "mulw" then
+ (match (string_drop _s2101_ (projT1 (string_length "mulw"))) with
+ | _s2102_ =>
+ (spc_matches_prefix _s2102_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1584_ _)) =>
- (match (string_drop _s1583_ _s1584_) with
- | _s1585_ =>
- (reg_name_matches_prefix _s1585_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2103_ _) =>
+ (match (string_drop _s2102_ _s2103_) with
+ | _s2104_ =>
+ (reg_name_matches_prefix _s2104_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1586_ _)) =>
- (match (string_drop _s1585_ _s1586_) with
- | _s1587_ =>
- (sep_matches_prefix _s1587_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s2105_ _) =>
+ (match (string_drop _s2104_ _s2105_) with
+ | _s2106_ =>
+ (sep_matches_prefix _s2106_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1588_ _)) =>
- (match (string_drop _s1587_ _s1588_) with
- | _s1589_ =>
- (reg_name_matches_prefix _s1589_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2107_ _) =>
+ (match (string_drop _s2106_ _s2107_) with
+ | _s2108_ =>
+ (reg_name_matches_prefix _s2108_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s1590_ _)) =>
- (match (string_drop _s1589_ _s1590_) with
- | _s1591_ =>
- (sep_matches_prefix _s1591_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s2109_ _) =>
+ (match (string_drop _s2108_ _s2109_) with
+ | _s2110_ =>
+ (sep_matches_prefix _s2110_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s1592_ _)) =>
- (match (string_drop _s1591_ _s1592_) with
- | _s1593_ =>
- (reg_name_matches_prefix _s1593_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2111_ _) =>
+ (match (string_drop _s2110_ _s2111_) with
+ | _s2112_ =>
+ (reg_name_matches_prefix _s2112_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((rs2, existT _ _s1594_ _)) =>
- let p0_ :=
- string_drop _s1593_ _s1594_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 5)))
+ returnm (match w__5 with
+ | Some (rs2, existT _ _s2113_ _) =>
+ let p0_ := string_drop _s2112_ _s2113_ in
+ if generic_eq p0_ "" then
+ Some (rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 5))).
-Definition _s1562_ (_s1563_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s1564_ := _s1563_ in
- (if ((string_startswith _s1564_ "rem")) then
- (match (string_drop _s1564_ (projT1 (string_length "rem"))) with
- | _s1565_ =>
- (maybe_not_u_matches_prefix _s1565_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s2081_ (_s2082_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5))) :=
+ let _s2083_ := _s2082_ in
+ (if string_startswith _s2083_ "rem" then
+ (match (string_drop _s2083_ (projT1 (string_length "rem"))) with
+ | _s2084_ =>
+ (maybe_not_u_matches_prefix _s2084_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s1566_ _)) =>
- (match (string_drop _s1565_ _s1566_) with
- | _s1567_ =>
- (spc_matches_prefix _s1567_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s2085_ _) =>
+ (match (string_drop _s2084_ _s2085_) with
+ | _s2086_ =>
+ (spc_matches_prefix _s2086_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1568_ _)) =>
- (match (string_drop _s1567_ _s1568_) with
- | _s1569_ =>
- (reg_name_matches_prefix _s1569_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2087_ _) =>
+ (match (string_drop _s2086_ _s2087_) with
+ | _s2088_ =>
+ (reg_name_matches_prefix _s2088_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1570_ _)) =>
- (match (string_drop _s1569_ _s1570_) with
- | _s1571_ =>
- (sep_matches_prefix _s1571_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s2089_ _) =>
+ (match (string_drop _s2088_ _s2089_) with
+ | _s2090_ =>
+ (sep_matches_prefix _s2090_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1572_ _)) =>
- (match (string_drop _s1571_ _s1572_) with
- | _s1573_ =>
- (reg_name_matches_prefix _s1573_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2091_ _) =>
+ (match (string_drop _s2090_ _s2091_) with
+ | _s2092_ =>
+ (reg_name_matches_prefix _s2092_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1574_ _)) =>
- (match (string_drop _s1573_ _s1574_) with
- | _s1575_ =>
- (sep_matches_prefix _s1575_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s2093_ _) =>
+ (match (string_drop _s2092_ _s2093_) with
+ | _s2094_ =>
+ (sep_matches_prefix _s2094_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1576_ _)) =>
- (match (string_drop _s1575_ _s1576_) with
- | _s1577_ =>
- (reg_name_matches_prefix _s1577_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2095_ _) =>
+ (match (string_drop _s2094_ _s2095_) with
+ | _s2096_ =>
+ (reg_name_matches_prefix _s2096_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s1578_ _)) =>
- let p0_ :=
- string_drop _s1577_ _s1578_ in
- if ((generic_eq p0_ "")) then
- Some
- ((s, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2097_ _) =>
+ let p0_ :=
+ string_drop _s2096_ _s2097_ in
+ if generic_eq p0_ "" then
+ Some (s, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5))).
-Definition _s1544_ (_s1545_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s1546_ := _s1545_ in
- (if ((string_startswith _s1546_ "div")) then
- (match (string_drop _s1546_ (projT1 (string_length "div"))) with
- | _s1547_ =>
- (maybe_not_u_matches_prefix _s1547_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s2063_ (_s2064_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5))) :=
+ let _s2065_ := _s2064_ in
+ (if string_startswith _s2065_ "div" then
+ (match (string_drop _s2065_ (projT1 (string_length "div"))) with
+ | _s2066_ =>
+ (maybe_not_u_matches_prefix _s2066_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s1548_ _)) =>
- (match (string_drop _s1547_ _s1548_) with
- | _s1549_ =>
- (spc_matches_prefix _s1549_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s2067_ _) =>
+ (match (string_drop _s2066_ _s2067_) with
+ | _s2068_ =>
+ (spc_matches_prefix _s2068_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1550_ _)) =>
- (match (string_drop _s1549_ _s1550_) with
- | _s1551_ =>
- (reg_name_matches_prefix _s1551_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2069_ _) =>
+ (match (string_drop _s2068_ _s2069_) with
+ | _s2070_ =>
+ (reg_name_matches_prefix _s2070_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1552_ _)) =>
- (match (string_drop _s1551_ _s1552_) with
- | _s1553_ =>
- (sep_matches_prefix _s1553_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s2071_ _) =>
+ (match (string_drop _s2070_ _s2071_) with
+ | _s2072_ =>
+ (sep_matches_prefix _s2072_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1554_ _)) =>
- (match (string_drop _s1553_ _s1554_) with
- | _s1555_ =>
- (reg_name_matches_prefix _s1555_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2073_ _) =>
+ (match (string_drop _s2072_ _s2073_) with
+ | _s2074_ =>
+ (reg_name_matches_prefix _s2074_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1556_ _)) =>
- (match (string_drop _s1555_ _s1556_) with
- | _s1557_ =>
- (sep_matches_prefix _s1557_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s2075_ _) =>
+ (match (string_drop _s2074_ _s2075_) with
+ | _s2076_ =>
+ (sep_matches_prefix _s2076_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1558_ _)) =>
- (match (string_drop _s1557_ _s1558_) with
- | _s1559_ =>
- (reg_name_matches_prefix _s1559_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2077_ _) =>
+ (match (string_drop _s2076_ _s2077_) with
+ | _s2078_ =>
+ (reg_name_matches_prefix _s2078_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s1560_ _)) =>
- let p0_ :=
- string_drop _s1559_ _s1560_ in
- if ((generic_eq p0_ "")) then
- Some
- ((s, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2079_ _) =>
+ let p0_ :=
+ string_drop _s2078_ _s2079_ in
+ if generic_eq p0_ "" then
+ Some (s, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5))).
-Definition _s1527_ (_s1528_ : string)
+Definition _s2046_ (_s2047_ : string)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1528_ with
- | _s1529_ =>
- (mul_mnemonic_matches_prefix _s1529_) >>= fun w__0 : option (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ (match _s2047_ with
+ | _s2048_ =>
+ (mul_mnemonic_matches_prefix _s2048_) >>= fun w__0 : option (((bool * bool * bool) * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some (((high, signed1, signed2), existT _ _s1530_ _)) =>
- (match (string_drop _s1529_ _s1530_) with
- | _s1531_ =>
- (spc_matches_prefix _s1531_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some ((high, signed1, signed2), existT _ _s2049_ _) =>
+ (match (string_drop _s2048_ _s2049_) with
+ | _s2050_ =>
+ (spc_matches_prefix _s2050_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1532_ _)) =>
- (match (string_drop _s1531_ _s1532_) with
- | _s1533_ =>
- (reg_name_matches_prefix _s1533_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2051_ _) =>
+ (match (string_drop _s2050_ _s2051_) with
+ | _s2052_ =>
+ (reg_name_matches_prefix _s2052_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1534_ _)) =>
- (match (string_drop _s1533_ _s1534_) with
- | _s1535_ =>
- (sep_matches_prefix _s1535_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s2053_ _) =>
+ (match (string_drop _s2052_ _s2053_) with
+ | _s2054_ =>
+ (sep_matches_prefix _s2054_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1536_ _)) =>
- (match (string_drop _s1535_ _s1536_) with
- | _s1537_ =>
- (reg_name_matches_prefix _s1537_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2055_ _) =>
+ (match (string_drop _s2054_ _s2055_) with
+ | _s2056_ =>
+ (reg_name_matches_prefix _s2056_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1538_ _)) =>
- (match (string_drop _s1537_ _s1538_) with
- | _s1539_ =>
- (sep_matches_prefix _s1539_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s2057_ _) =>
+ (match (string_drop _s2056_ _s2057_) with
+ | _s2058_ =>
+ (sep_matches_prefix _s2058_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1540_ _)) =>
- (match (string_drop _s1539_ _s1540_) with
- | _s1541_ =>
- (reg_name_matches_prefix _s1541_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2059_ _) =>
+ (match (string_drop _s2058_ _s2059_) with
+ | _s2060_ =>
+ (reg_name_matches_prefix _s2060_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s1542_ _)) =>
- let p0_ :=
- string_drop _s1541_ _s1542_ in
- if ((generic_eq p0_ "")) then
- Some
- ((high, signed1, signed2, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s2061_ _) =>
+ let p0_ :=
+ string_drop _s2060_ _s2061_ in
+ if generic_eq p0_ "" then
+ Some
+ (high, signed1, signed2, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5))).
-Definition _s1515_ (_s1516_ : string)
-: M (option ((mword 5 * mword 5))) :=
-
- let _s1517_ := _s1516_ in
- (if ((string_startswith _s1517_ "c.add")) then
- (match (string_drop _s1517_ (projT1 (string_length "c.add"))) with
- | _s1518_ =>
- (spc_matches_prefix _s1518_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s2034_ (_s2035_ : string) : M (option ((mword 5 * mword 5))) :=
+ let _s2036_ := _s2035_ in
+ (if string_startswith _s2036_ "c.add" then
+ (match (string_drop _s2036_ (projT1 (string_length "c.add"))) with
+ | _s2037_ =>
+ (spc_matches_prefix _s2037_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1519_ _)) =>
- (match (string_drop _s1518_ _s1519_) with
- | _s1520_ =>
- (reg_name_matches_prefix _s1520_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2038_ _) =>
+ (match (string_drop _s2037_ _s2038_) with
+ | _s2039_ =>
+ (reg_name_matches_prefix _s2039_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1521_ _)) =>
- (match (string_drop _s1520_ _s1521_) with
- | _s1522_ =>
- (sep_matches_prefix _s1522_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s2040_ _) =>
+ (match (string_drop _s2039_ _s2040_) with
+ | _s2041_ =>
+ (sep_matches_prefix _s2041_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1523_ _)) =>
- (match (string_drop _s1522_ _s1523_) with
- | _s1524_ =>
- (reg_name_matches_prefix _s1524_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2042_ _) =>
+ (match (string_drop _s2041_ _s2042_) with
+ | _s2043_ =>
+ (reg_name_matches_prefix _s2043_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1525_ _)) =>
- let p0_ := string_drop _s1524_ _s1525_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 5 * mword 5)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s2044_ _) =>
+ let p0_ := string_drop _s2043_ _s2044_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- else returnm (None : option ((mword 5 * mword 5))))
+ else returnm None)
: M (option ((mword 5 * mword 5))).
-Definition _s1503_ (_s1504_ : string)
-: M (option ((mword 5 * mword 5))) :=
-
- let _s1505_ := _s1504_ in
- (if ((string_startswith _s1505_ "c.mv")) then
- (match (string_drop _s1505_ (projT1 (string_length "c.mv"))) with
- | _s1506_ =>
- (spc_matches_prefix _s1506_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s2022_ (_s2023_ : string) : M (option ((mword 5 * mword 5))) :=
+ let _s2024_ := _s2023_ in
+ (if string_startswith _s2024_ "c.mv" then
+ (match (string_drop _s2024_ (projT1 (string_length "c.mv"))) with
+ | _s2025_ =>
+ (spc_matches_prefix _s2025_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1507_ _)) =>
- (match (string_drop _s1506_ _s1507_) with
- | _s1508_ =>
- (reg_name_matches_prefix _s1508_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2026_ _) =>
+ (match (string_drop _s2025_ _s2026_) with
+ | _s2027_ =>
+ (reg_name_matches_prefix _s2027_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1509_ _)) =>
- (match (string_drop _s1508_ _s1509_) with
- | _s1510_ =>
- (sep_matches_prefix _s1510_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s2028_ _) =>
+ (match (string_drop _s2027_ _s2028_) with
+ | _s2029_ =>
+ (sep_matches_prefix _s2029_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1511_ _)) =>
- (match (string_drop _s1510_ _s1511_) with
- | _s1512_ =>
- (reg_name_matches_prefix _s1512_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2030_ _) =>
+ (match (string_drop _s2029_ _s2030_) with
+ | _s2031_ =>
+ (reg_name_matches_prefix _s2031_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1513_ _)) =>
- let p0_ := string_drop _s1512_ _s1513_ in
- if ((generic_eq p0_ "")) then Some ((rd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 5 * mword 5)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s2032_ _) =>
+ let p0_ := string_drop _s2031_ _s2032_ in
+ if generic_eq p0_ "" then Some (rd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- else returnm (None : option ((mword 5 * mword 5))))
+ else returnm None)
: M (option ((mword 5 * mword 5))).
-Definition _s1495_ (_s1496_ : string)
-: M (option (mword 5)) :=
-
- let _s1497_ := _s1496_ in
- (if ((string_startswith _s1497_ "c.jalr")) then
- (match (string_drop _s1497_ (projT1 (string_length "c.jalr"))) with
- | _s1498_ =>
- (spc_matches_prefix _s1498_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s2014_ (_s2015_ : string) : M (option (mword 5)) :=
+ let _s2016_ := _s2015_ in
+ (if string_startswith _s2016_ "c.jalr" then
+ (match (string_drop _s2016_ (projT1 (string_length "c.jalr"))) with
+ | _s2017_ =>
+ (spc_matches_prefix _s2017_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1499_ _)) =>
- (match (string_drop _s1498_ _s1499_) with
- | _s1500_ =>
- (reg_name_matches_prefix _s1500_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2018_ _) =>
+ (match (string_drop _s2017_ _s2018_) with
+ | _s2019_ =>
+ (reg_name_matches_prefix _s2019_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((rs1, existT _ _s1501_ _)) =>
- let p0_ := string_drop _s1500_ _s1501_ in
- if ((generic_eq p0_ "")) then Some (rs1)
- else None
- | _ => None
- end)
- : option (mword 5))
+ returnm (match w__1 with
+ | Some (rs1, existT _ _s2020_ _) =>
+ let p0_ := string_drop _s2019_ _s2020_ in
+ if generic_eq p0_ "" then Some rs1
+ else None
+ | _ => None
+ end)
end)
: M (option (mword 5))
- | _ => returnm (None : option (mword 5))
+ | _ => returnm None
end)
: M (option (mword 5))
end)
: M (option (mword 5))
- else returnm (None : option (mword 5)))
+ else returnm None)
: M (option (mword 5)).
-Definition _s1487_ (_s1488_ : string)
-: M (option (mword 5)) :=
-
- let _s1489_ := _s1488_ in
- (if ((string_startswith _s1489_ "c.jr")) then
- (match (string_drop _s1489_ (projT1 (string_length "c.jr"))) with
- | _s1490_ =>
- (spc_matches_prefix _s1490_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s2006_ (_s2007_ : string) : M (option (mword 5)) :=
+ let _s2008_ := _s2007_ in
+ (if string_startswith _s2008_ "c.jr" then
+ (match (string_drop _s2008_ (projT1 (string_length "c.jr"))) with
+ | _s2009_ =>
+ (spc_matches_prefix _s2009_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1491_ _)) =>
- (match (string_drop _s1490_ _s1491_) with
- | _s1492_ =>
- (reg_name_matches_prefix _s1492_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s2010_ _) =>
+ (match (string_drop _s2009_ _s2010_) with
+ | _s2011_ =>
+ (reg_name_matches_prefix _s2011_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((rs1, existT _ _s1493_ _)) =>
- let p0_ := string_drop _s1492_ _s1493_ in
- if ((generic_eq p0_ "")) then Some (rs1)
- else None
- | _ => None
- end)
- : option (mword 5))
+ returnm (match w__1 with
+ | Some (rs1, existT _ _s2012_ _) =>
+ let p0_ := string_drop _s2011_ _s2012_ in
+ if generic_eq p0_ "" then Some rs1
+ else None
+ | _ => None
+ end)
end)
: M (option (mword 5))
- | _ => returnm (None : option (mword 5))
+ | _ => returnm None
end)
: M (option (mword 5))
end)
: M (option (mword 5))
- else returnm (None : option (mword 5)))
+ else returnm None)
: M (option (mword 5)).
-Definition _s1475_ (_s1476_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1477_ := _s1476_ in
- (if ((string_startswith _s1477_ "c.sdsp")) then
- (match (string_drop _s1477_ (projT1 (string_length "c.sdsp"))) with
- | _s1478_ =>
- (spc_matches_prefix _s1478_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1994_ (_s1995_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1996_ := _s1995_ in
+ (if string_startswith _s1996_ "c.sdsp" then
+ (match (string_drop _s1996_ (projT1 (string_length "c.sdsp"))) with
+ | _s1997_ =>
+ (spc_matches_prefix _s1997_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1479_ _)) =>
- (match (string_drop _s1478_ _s1479_) with
- | _s1480_ =>
- (reg_name_matches_prefix _s1480_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1998_ _) =>
+ (match (string_drop _s1997_ _s1998_) with
+ | _s1999_ =>
+ (reg_name_matches_prefix _s1999_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs2, existT _ _s1481_ _)) =>
- (match (string_drop _s1480_ _s1481_) with
- | _s1482_ =>
- (sep_matches_prefix _s1482_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s2000_ _) =>
+ (match (string_drop _s1999_ _s2000_) with
+ | _s2001_ =>
+ (sep_matches_prefix _s2001_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1483_ _)) =>
- match (string_drop _s1482_ _s1483_) with
- | _s1484_ =>
- match (hex_bits_6_matches_prefix _s1484_) with
- | Some ((uimm, existT _ _s1485_ _)) =>
- let p0_ := string_drop _s1484_ _s1485_ in
- if ((generic_eq p0_ "")) then Some ((rs2, uimm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s2002_ _) =>
+ match (string_drop _s2001_ _s2002_) with
+ | _s2003_ =>
+ match (hex_bits_6_matches_prefix _s2003_) with
+ | Some (uimm, existT _ _s2004_ _) =>
+ let p0_ := string_drop _s2003_ _s2004_ in
+ if generic_eq p0_ "" then Some (rs2, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1463_ (_s1464_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1465_ := _s1464_ in
- (if ((string_startswith _s1465_ "c.swsp")) then
- (match (string_drop _s1465_ (projT1 (string_length "c.swsp"))) with
- | _s1466_ =>
- (spc_matches_prefix _s1466_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1982_ (_s1983_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1984_ := _s1983_ in
+ (if string_startswith _s1984_ "c.swsp" then
+ (match (string_drop _s1984_ (projT1 (string_length "c.swsp"))) with
+ | _s1985_ =>
+ (spc_matches_prefix _s1985_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1467_ _)) =>
- (match (string_drop _s1466_ _s1467_) with
- | _s1468_ =>
- (reg_name_matches_prefix _s1468_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1986_ _) =>
+ (match (string_drop _s1985_ _s1986_) with
+ | _s1987_ =>
+ (reg_name_matches_prefix _s1987_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1469_ _)) =>
- (match (string_drop _s1468_ _s1469_) with
- | _s1470_ =>
- (sep_matches_prefix _s1470_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1988_ _) =>
+ (match (string_drop _s1987_ _s1988_) with
+ | _s1989_ =>
+ (sep_matches_prefix _s1989_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1471_ _)) =>
- match (string_drop _s1470_ _s1471_) with
- | _s1472_ =>
- match (hex_bits_6_matches_prefix _s1472_) with
- | Some ((uimm, existT _ _s1473_ _)) =>
- let p0_ := string_drop _s1472_ _s1473_ in
- if ((generic_eq p0_ "")) then Some ((rd, uimm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1990_ _) =>
+ match (string_drop _s1989_ _s1990_) with
+ | _s1991_ =>
+ match (hex_bits_6_matches_prefix _s1991_) with
+ | Some (uimm, existT _ _s1992_ _) =>
+ let p0_ := string_drop _s1991_ _s1992_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1451_ (_s1452_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1453_ := _s1452_ in
- (if ((string_startswith _s1453_ "c.ldsp")) then
- (match (string_drop _s1453_ (projT1 (string_length "c.ldsp"))) with
- | _s1454_ =>
- (spc_matches_prefix _s1454_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1970_ (_s1971_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1972_ := _s1971_ in
+ (if string_startswith _s1972_ "c.ldsp" then
+ (match (string_drop _s1972_ (projT1 (string_length "c.ldsp"))) with
+ | _s1973_ =>
+ (spc_matches_prefix _s1973_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1455_ _)) =>
- (match (string_drop _s1454_ _s1455_) with
- | _s1456_ =>
- (reg_name_matches_prefix _s1456_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1974_ _) =>
+ (match (string_drop _s1973_ _s1974_) with
+ | _s1975_ =>
+ (reg_name_matches_prefix _s1975_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1457_ _)) =>
- (match (string_drop _s1456_ _s1457_) with
- | _s1458_ =>
- (sep_matches_prefix _s1458_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1976_ _) =>
+ (match (string_drop _s1975_ _s1976_) with
+ | _s1977_ =>
+ (sep_matches_prefix _s1977_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1459_ _)) =>
- match (string_drop _s1458_ _s1459_) with
- | _s1460_ =>
- match (hex_bits_6_matches_prefix _s1460_) with
- | Some ((uimm, existT _ _s1461_ _)) =>
- let p0_ := string_drop _s1460_ _s1461_ in
- if ((generic_eq p0_ "")) then Some ((rd, uimm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1978_ _) =>
+ match (string_drop _s1977_ _s1978_) with
+ | _s1979_ =>
+ match (hex_bits_6_matches_prefix _s1979_) with
+ | Some (uimm, existT _ _s1980_ _) =>
+ let p0_ := string_drop _s1979_ _s1980_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1439_ (_s1440_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1441_ := _s1440_ in
- (if ((string_startswith _s1441_ "c.lwsp")) then
- (match (string_drop _s1441_ (projT1 (string_length "c.lwsp"))) with
- | _s1442_ =>
- (spc_matches_prefix _s1442_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1958_ (_s1959_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1960_ := _s1959_ in
+ (if string_startswith _s1960_ "c.lwsp" then
+ (match (string_drop _s1960_ (projT1 (string_length "c.lwsp"))) with
+ | _s1961_ =>
+ (spc_matches_prefix _s1961_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1443_ _)) =>
- (match (string_drop _s1442_ _s1443_) with
- | _s1444_ =>
- (reg_name_matches_prefix _s1444_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1962_ _) =>
+ (match (string_drop _s1961_ _s1962_) with
+ | _s1963_ =>
+ (reg_name_matches_prefix _s1963_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1445_ _)) =>
- (match (string_drop _s1444_ _s1445_) with
- | _s1446_ =>
- (sep_matches_prefix _s1446_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1964_ _) =>
+ (match (string_drop _s1963_ _s1964_) with
+ | _s1965_ =>
+ (sep_matches_prefix _s1965_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1447_ _)) =>
- match (string_drop _s1446_ _s1447_) with
- | _s1448_ =>
- match (hex_bits_6_matches_prefix _s1448_) with
- | Some ((uimm, existT _ _s1449_ _)) =>
- let p0_ := string_drop _s1448_ _s1449_ in
- if ((generic_eq p0_ "")) then Some ((rd, uimm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1966_ _) =>
+ match (string_drop _s1965_ _s1966_) with
+ | _s1967_ =>
+ match (hex_bits_6_matches_prefix _s1967_) with
+ | Some (uimm, existT _ _s1968_ _) =>
+ let p0_ := string_drop _s1967_ _s1968_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1427_ (_s1428_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1429_ := _s1428_ in
- (if ((string_startswith _s1429_ "c.slli")) then
- (match (string_drop _s1429_ (projT1 (string_length "c.slli"))) with
- | _s1430_ =>
- (spc_matches_prefix _s1430_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1946_ (_s1947_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1948_ := _s1947_ in
+ (if string_startswith _s1948_ "c.slli" then
+ (match (string_drop _s1948_ (projT1 (string_length "c.slli"))) with
+ | _s1949_ =>
+ (spc_matches_prefix _s1949_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1431_ _)) =>
- (match (string_drop _s1430_ _s1431_) with
- | _s1432_ =>
- (reg_name_matches_prefix _s1432_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1950_ _) =>
+ (match (string_drop _s1949_ _s1950_) with
+ | _s1951_ =>
+ (reg_name_matches_prefix _s1951_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1433_ _)) =>
- (match (string_drop _s1432_ _s1433_) with
- | _s1434_ =>
- (sep_matches_prefix _s1434_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1952_ _) =>
+ (match (string_drop _s1951_ _s1952_) with
+ | _s1953_ =>
+ (sep_matches_prefix _s1953_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1435_ _)) =>
- match (string_drop _s1434_ _s1435_) with
- | _s1436_ =>
- match (hex_bits_6_matches_prefix _s1436_) with
- | Some ((shamt, existT _ _s1437_ _)) =>
- let p0_ := string_drop _s1436_ _s1437_ in
- if ((generic_eq p0_ "")) then Some ((rsd, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1954_ _) =>
+ match (string_drop _s1953_ _s1954_) with
+ | _s1955_ =>
+ match (hex_bits_6_matches_prefix _s1955_) with
+ | Some (shamt, existT _ _s1956_ _) =>
+ let p0_ := string_drop _s1955_ _s1956_ in
+ if generic_eq p0_ "" then Some (rsd, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1415_ (_s1416_ : string)
-: M (option ((mword 3 * mword 8))) :=
-
- let _s1417_ := _s1416_ in
- (if ((string_startswith _s1417_ "c.bnez")) then
- (match (string_drop _s1417_ (projT1 (string_length "c.bnez"))) with
- | _s1418_ =>
- (spc_matches_prefix _s1418_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1934_ (_s1935_ : string) : M (option ((mword 3 * mword 8))) :=
+ let _s1936_ := _s1935_ in
+ (if string_startswith _s1936_ "c.bnez" then
+ (match (string_drop _s1936_ (projT1 (string_length "c.bnez"))) with
+ | _s1937_ =>
+ (spc_matches_prefix _s1937_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1419_ _)) =>
- (match (string_drop _s1418_ _s1419_) with
- | _s1420_ =>
- (creg_name_matches_prefix _s1420_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1938_ _) =>
+ (match (string_drop _s1937_ _s1938_) with
+ | _s1939_ =>
+ (creg_name_matches_prefix _s1939_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs, existT _ _s1421_ _)) =>
- (match (string_drop _s1420_ _s1421_) with
- | _s1422_ =>
- (sep_matches_prefix _s1422_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs, existT _ _s1940_ _) =>
+ (match (string_drop _s1939_ _s1940_) with
+ | _s1941_ =>
+ (sep_matches_prefix _s1941_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1423_ _)) =>
- match (string_drop _s1422_ _s1423_) with
- | _s1424_ =>
- match (hex_bits_8_matches_prefix _s1424_) with
- | Some ((imm, existT _ _s1425_ _)) =>
- let p0_ := string_drop _s1424_ _s1425_ in
- if ((generic_eq p0_ "")) then Some ((rs, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1942_ _) =>
+ match (string_drop _s1941_ _s1942_) with
+ | _s1943_ =>
+ match (hex_bits_8_matches_prefix _s1943_) with
+ | Some (imm, existT _ _s1944_ _) =>
+ let p0_ := string_drop _s1943_ _s1944_ in
+ if generic_eq p0_ "" then Some (rs, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- else returnm (None : option ((mword 3 * mword 8))))
+ else returnm None)
: M (option ((mword 3 * mword 8))).
-Definition _s1403_ (_s1404_ : string)
-: M (option ((mword 3 * mword 8))) :=
-
- let _s1405_ := _s1404_ in
- (if ((string_startswith _s1405_ "c.beqz")) then
- (match (string_drop _s1405_ (projT1 (string_length "c.beqz"))) with
- | _s1406_ =>
- (spc_matches_prefix _s1406_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1922_ (_s1923_ : string) : M (option ((mword 3 * mword 8))) :=
+ let _s1924_ := _s1923_ in
+ (if string_startswith _s1924_ "c.beqz" then
+ (match (string_drop _s1924_ (projT1 (string_length "c.beqz"))) with
+ | _s1925_ =>
+ (spc_matches_prefix _s1925_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1407_ _)) =>
- (match (string_drop _s1406_ _s1407_) with
- | _s1408_ =>
- (creg_name_matches_prefix _s1408_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1926_ _) =>
+ (match (string_drop _s1925_ _s1926_) with
+ | _s1927_ =>
+ (creg_name_matches_prefix _s1927_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs, existT _ _s1409_ _)) =>
- (match (string_drop _s1408_ _s1409_) with
- | _s1410_ =>
- (sep_matches_prefix _s1410_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs, existT _ _s1928_ _) =>
+ (match (string_drop _s1927_ _s1928_) with
+ | _s1929_ =>
+ (sep_matches_prefix _s1929_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1411_ _)) =>
- match (string_drop _s1410_ _s1411_) with
- | _s1412_ =>
- match (hex_bits_8_matches_prefix _s1412_) with
- | Some ((imm, existT _ _s1413_ _)) =>
- let p0_ := string_drop _s1412_ _s1413_ in
- if ((generic_eq p0_ "")) then Some ((rs, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1930_ _) =>
+ match (string_drop _s1929_ _s1930_) with
+ | _s1931_ =>
+ match (hex_bits_8_matches_prefix _s1931_) with
+ | Some (imm, existT _ _s1932_ _) =>
+ let p0_ := string_drop _s1931_ _s1932_ in
+ if generic_eq p0_ "" then Some (rs, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- else returnm (None : option ((mword 3 * mword 8))))
+ else returnm None)
: M (option ((mword 3 * mword 8))).
-Definition _s1395_ (_s1396_ : string)
-: M (option (mword 11)) :=
-
- let _s1397_ := _s1396_ in
- (if ((string_startswith _s1397_ "c.j")) then
- (match (string_drop _s1397_ (projT1 (string_length "c.j"))) with
- | _s1398_ =>
- (spc_matches_prefix _s1398_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s1399_ _)) =>
- match (string_drop _s1398_ _s1399_) with
- | _s1400_ =>
- match (hex_bits_11_matches_prefix _s1400_) with
- | Some ((imm, existT _ _s1401_ _)) =>
- let p0_ := string_drop _s1400_ _s1401_ in
- if ((generic_eq p0_ "")) then Some (imm)
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 11))
+Definition _s1914_ (_s1915_ : string) : M (option (mword 11)) :=
+ let _s1916_ := _s1915_ in
+ (if string_startswith _s1916_ "c.j" then
+ (match (string_drop _s1916_ (projT1 (string_length "c.j"))) with
+ | _s1917_ =>
+ (spc_matches_prefix _s1917_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s1918_ _) =>
+ match (string_drop _s1917_ _s1918_) with
+ | _s1919_ =>
+ match (hex_bits_11_matches_prefix _s1919_) with
+ | Some (imm, existT _ _s1920_ _) =>
+ let p0_ := string_drop _s1919_ _s1920_ in
+ if generic_eq p0_ "" then Some imm
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 11))
- else returnm (None : option (mword 11)))
+ else returnm None)
: M (option (mword 11)).
-Definition _s1383_ (_s1384_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s1385_ := _s1384_ in
- (if ((string_startswith _s1385_ "c.addw")) then
- (match (string_drop _s1385_ (projT1 (string_length "c.addw"))) with
- | _s1386_ =>
- (spc_matches_prefix _s1386_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1902_ (_s1903_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s1904_ := _s1903_ in
+ (if string_startswith _s1904_ "c.addw" then
+ (match (string_drop _s1904_ (projT1 (string_length "c.addw"))) with
+ | _s1905_ =>
+ (spc_matches_prefix _s1905_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1387_ _)) =>
- (match (string_drop _s1386_ _s1387_) with
- | _s1388_ =>
- (creg_name_matches_prefix _s1388_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1906_ _) =>
+ (match (string_drop _s1905_ _s1906_) with
+ | _s1907_ =>
+ (creg_name_matches_prefix _s1907_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1389_ _)) =>
- (match (string_drop _s1388_ _s1389_) with
- | _s1390_ =>
- (sep_matches_prefix _s1390_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1908_ _) =>
+ (match (string_drop _s1907_ _s1908_) with
+ | _s1909_ =>
+ (sep_matches_prefix _s1909_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1391_ _)) =>
- (match (string_drop _s1390_ _s1391_) with
- | _s1392_ =>
- (creg_name_matches_prefix _s1392_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1910_ _) =>
+ (match (string_drop _s1909_ _s1910_) with
+ | _s1911_ =>
+ (creg_name_matches_prefix _s1911_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1393_ _)) =>
- let p0_ := string_drop _s1392_ _s1393_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s1912_ _) =>
+ let p0_ := string_drop _s1911_ _s1912_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s1371_ (_s1372_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s1373_ := _s1372_ in
- (if ((string_startswith _s1373_ "c.subw")) then
- (match (string_drop _s1373_ (projT1 (string_length "c.subw"))) with
- | _s1374_ =>
- (spc_matches_prefix _s1374_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1890_ (_s1891_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s1892_ := _s1891_ in
+ (if string_startswith _s1892_ "c.subw" then
+ (match (string_drop _s1892_ (projT1 (string_length "c.subw"))) with
+ | _s1893_ =>
+ (spc_matches_prefix _s1893_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1375_ _)) =>
- (match (string_drop _s1374_ _s1375_) with
- | _s1376_ =>
- (creg_name_matches_prefix _s1376_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1894_ _) =>
+ (match (string_drop _s1893_ _s1894_) with
+ | _s1895_ =>
+ (creg_name_matches_prefix _s1895_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1377_ _)) =>
- (match (string_drop _s1376_ _s1377_) with
- | _s1378_ =>
- (sep_matches_prefix _s1378_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1896_ _) =>
+ (match (string_drop _s1895_ _s1896_) with
+ | _s1897_ =>
+ (sep_matches_prefix _s1897_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1379_ _)) =>
- (match (string_drop _s1378_ _s1379_) with
- | _s1380_ =>
- (creg_name_matches_prefix _s1380_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1898_ _) =>
+ (match (string_drop _s1897_ _s1898_) with
+ | _s1899_ =>
+ (creg_name_matches_prefix _s1899_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1381_ _)) =>
- let p0_ := string_drop _s1380_ _s1381_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s1900_ _) =>
+ let p0_ := string_drop _s1899_ _s1900_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s1359_ (_s1360_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s1361_ := _s1360_ in
- (if ((string_startswith _s1361_ "c.and")) then
- (match (string_drop _s1361_ (projT1 (string_length "c.and"))) with
- | _s1362_ =>
- (spc_matches_prefix _s1362_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1878_ (_s1879_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s1880_ := _s1879_ in
+ (if string_startswith _s1880_ "c.and" then
+ (match (string_drop _s1880_ (projT1 (string_length "c.and"))) with
+ | _s1881_ =>
+ (spc_matches_prefix _s1881_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1363_ _)) =>
- (match (string_drop _s1362_ _s1363_) with
- | _s1364_ =>
- (creg_name_matches_prefix _s1364_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1882_ _) =>
+ (match (string_drop _s1881_ _s1882_) with
+ | _s1883_ =>
+ (creg_name_matches_prefix _s1883_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1365_ _)) =>
- (match (string_drop _s1364_ _s1365_) with
- | _s1366_ =>
- (sep_matches_prefix _s1366_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1884_ _) =>
+ (match (string_drop _s1883_ _s1884_) with
+ | _s1885_ =>
+ (sep_matches_prefix _s1885_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1367_ _)) =>
- (match (string_drop _s1366_ _s1367_) with
- | _s1368_ =>
- (creg_name_matches_prefix _s1368_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1886_ _) =>
+ (match (string_drop _s1885_ _s1886_) with
+ | _s1887_ =>
+ (creg_name_matches_prefix _s1887_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1369_ _)) =>
- let p0_ := string_drop _s1368_ _s1369_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s1888_ _) =>
+ let p0_ := string_drop _s1887_ _s1888_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s1347_ (_s1348_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s1349_ := _s1348_ in
- (if ((string_startswith _s1349_ "c.or")) then
- (match (string_drop _s1349_ (projT1 (string_length "c.or"))) with
- | _s1350_ =>
- (spc_matches_prefix _s1350_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1866_ (_s1867_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s1868_ := _s1867_ in
+ (if string_startswith _s1868_ "c.or" then
+ (match (string_drop _s1868_ (projT1 (string_length "c.or"))) with
+ | _s1869_ =>
+ (spc_matches_prefix _s1869_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1351_ _)) =>
- (match (string_drop _s1350_ _s1351_) with
- | _s1352_ =>
- (creg_name_matches_prefix _s1352_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1870_ _) =>
+ (match (string_drop _s1869_ _s1870_) with
+ | _s1871_ =>
+ (creg_name_matches_prefix _s1871_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1353_ _)) =>
- (match (string_drop _s1352_ _s1353_) with
- | _s1354_ =>
- (sep_matches_prefix _s1354_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1872_ _) =>
+ (match (string_drop _s1871_ _s1872_) with
+ | _s1873_ =>
+ (sep_matches_prefix _s1873_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1355_ _)) =>
- (match (string_drop _s1354_ _s1355_) with
- | _s1356_ =>
- (creg_name_matches_prefix _s1356_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1874_ _) =>
+ (match (string_drop _s1873_ _s1874_) with
+ | _s1875_ =>
+ (creg_name_matches_prefix _s1875_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1357_ _)) =>
- let p0_ := string_drop _s1356_ _s1357_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s1876_ _) =>
+ let p0_ := string_drop _s1875_ _s1876_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s1335_ (_s1336_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s1337_ := _s1336_ in
- (if ((string_startswith _s1337_ "c.xor")) then
- (match (string_drop _s1337_ (projT1 (string_length "c.xor"))) with
- | _s1338_ =>
- (spc_matches_prefix _s1338_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1854_ (_s1855_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s1856_ := _s1855_ in
+ (if string_startswith _s1856_ "c.xor" then
+ (match (string_drop _s1856_ (projT1 (string_length "c.xor"))) with
+ | _s1857_ =>
+ (spc_matches_prefix _s1857_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1339_ _)) =>
- (match (string_drop _s1338_ _s1339_) with
- | _s1340_ =>
- (creg_name_matches_prefix _s1340_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1858_ _) =>
+ (match (string_drop _s1857_ _s1858_) with
+ | _s1859_ =>
+ (creg_name_matches_prefix _s1859_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1341_ _)) =>
- (match (string_drop _s1340_ _s1341_) with
- | _s1342_ =>
- (sep_matches_prefix _s1342_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1860_ _) =>
+ (match (string_drop _s1859_ _s1860_) with
+ | _s1861_ =>
+ (sep_matches_prefix _s1861_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1343_ _)) =>
- (match (string_drop _s1342_ _s1343_) with
- | _s1344_ =>
- (creg_name_matches_prefix _s1344_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1862_ _) =>
+ (match (string_drop _s1861_ _s1862_) with
+ | _s1863_ =>
+ (creg_name_matches_prefix _s1863_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1345_ _)) =>
- let p0_ := string_drop _s1344_ _s1345_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s1864_ _) =>
+ let p0_ := string_drop _s1863_ _s1864_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s1323_ (_s1324_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s1325_ := _s1324_ in
- (if ((string_startswith _s1325_ "c.sub")) then
- (match (string_drop _s1325_ (projT1 (string_length "c.sub"))) with
- | _s1326_ =>
- (spc_matches_prefix _s1326_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1842_ (_s1843_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s1844_ := _s1843_ in
+ (if string_startswith _s1844_ "c.sub" then
+ (match (string_drop _s1844_ (projT1 (string_length "c.sub"))) with
+ | _s1845_ =>
+ (spc_matches_prefix _s1845_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1327_ _)) =>
- (match (string_drop _s1326_ _s1327_) with
- | _s1328_ =>
- (creg_name_matches_prefix _s1328_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1846_ _) =>
+ (match (string_drop _s1845_ _s1846_) with
+ | _s1847_ =>
+ (creg_name_matches_prefix _s1847_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1329_ _)) =>
- (match (string_drop _s1328_ _s1329_) with
- | _s1330_ =>
- (sep_matches_prefix _s1330_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1848_ _) =>
+ (match (string_drop _s1847_ _s1848_) with
+ | _s1849_ =>
+ (sep_matches_prefix _s1849_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1331_ _)) =>
- (match (string_drop _s1330_ _s1331_) with
- | _s1332_ =>
- (creg_name_matches_prefix _s1332_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1850_ _) =>
+ (match (string_drop _s1849_ _s1850_) with
+ | _s1851_ =>
+ (creg_name_matches_prefix _s1851_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1333_ _)) =>
- let p0_ := string_drop _s1332_ _s1333_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s1852_ _) =>
+ let p0_ := string_drop _s1851_ _s1852_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s1311_ (_s1312_ : string)
-: M (option ((mword 3 * mword 6))) :=
-
- let _s1313_ := _s1312_ in
- (if ((string_startswith _s1313_ "c.andi")) then
- (match (string_drop _s1313_ (projT1 (string_length "c.andi"))) with
- | _s1314_ =>
- (spc_matches_prefix _s1314_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1830_ (_s1831_ : string) : M (option ((mword 3 * mword 6))) :=
+ let _s1832_ := _s1831_ in
+ (if string_startswith _s1832_ "c.andi" then
+ (match (string_drop _s1832_ (projT1 (string_length "c.andi"))) with
+ | _s1833_ =>
+ (spc_matches_prefix _s1833_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1315_ _)) =>
- (match (string_drop _s1314_ _s1315_) with
- | _s1316_ =>
- (creg_name_matches_prefix _s1316_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1834_ _) =>
+ (match (string_drop _s1833_ _s1834_) with
+ | _s1835_ =>
+ (creg_name_matches_prefix _s1835_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1317_ _)) =>
- (match (string_drop _s1316_ _s1317_) with
- | _s1318_ =>
- (sep_matches_prefix _s1318_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1836_ _) =>
+ (match (string_drop _s1835_ _s1836_) with
+ | _s1837_ =>
+ (sep_matches_prefix _s1837_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1319_ _)) =>
- match (string_drop _s1318_ _s1319_) with
- | _s1320_ =>
- match (hex_bits_6_matches_prefix _s1320_) with
- | Some ((imm, existT _ _s1321_ _)) =>
- let p0_ := string_drop _s1320_ _s1321_ in
- if ((generic_eq p0_ "")) then Some ((rsd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1838_ _) =>
+ match (string_drop _s1837_ _s1838_) with
+ | _s1839_ =>
+ match (hex_bits_6_matches_prefix _s1839_) with
+ | Some (imm, existT _ _s1840_ _) =>
+ let p0_ := string_drop _s1839_ _s1840_ in
+ if generic_eq p0_ "" then Some (rsd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- else returnm (None : option ((mword 3 * mword 6))))
+ else returnm None)
: M (option ((mword 3 * mword 6))).
-Definition _s1299_ (_s1300_ : string)
-: M (option ((mword 3 * mword 6))) :=
-
- let _s1301_ := _s1300_ in
- (if ((string_startswith _s1301_ "c.srai")) then
- (match (string_drop _s1301_ (projT1 (string_length "c.srai"))) with
- | _s1302_ =>
- (spc_matches_prefix _s1302_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1818_ (_s1819_ : string) : M (option ((mword 3 * mword 6))) :=
+ let _s1820_ := _s1819_ in
+ (if string_startswith _s1820_ "c.srai" then
+ (match (string_drop _s1820_ (projT1 (string_length "c.srai"))) with
+ | _s1821_ =>
+ (spc_matches_prefix _s1821_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1303_ _)) =>
- (match (string_drop _s1302_ _s1303_) with
- | _s1304_ =>
- (creg_name_matches_prefix _s1304_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1822_ _) =>
+ (match (string_drop _s1821_ _s1822_) with
+ | _s1823_ =>
+ (creg_name_matches_prefix _s1823_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1305_ _)) =>
- (match (string_drop _s1304_ _s1305_) with
- | _s1306_ =>
- (sep_matches_prefix _s1306_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1824_ _) =>
+ (match (string_drop _s1823_ _s1824_) with
+ | _s1825_ =>
+ (sep_matches_prefix _s1825_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1307_ _)) =>
- match (string_drop _s1306_ _s1307_) with
- | _s1308_ =>
- match (hex_bits_6_matches_prefix _s1308_) with
- | Some ((shamt, existT _ _s1309_ _)) =>
- let p0_ := string_drop _s1308_ _s1309_ in
- if ((generic_eq p0_ "")) then Some ((rsd, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1826_ _) =>
+ match (string_drop _s1825_ _s1826_) with
+ | _s1827_ =>
+ match (hex_bits_6_matches_prefix _s1827_) with
+ | Some (shamt, existT _ _s1828_ _) =>
+ let p0_ := string_drop _s1827_ _s1828_ in
+ if generic_eq p0_ "" then Some (rsd, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- else returnm (None : option ((mword 3 * mword 6))))
+ else returnm None)
: M (option ((mword 3 * mword 6))).
-Definition _s1287_ (_s1288_ : string)
-: M (option ((mword 3 * mword 6))) :=
-
- let _s1289_ := _s1288_ in
- (if ((string_startswith _s1289_ "c.srli")) then
- (match (string_drop _s1289_ (projT1 (string_length "c.srli"))) with
- | _s1290_ =>
- (spc_matches_prefix _s1290_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1806_ (_s1807_ : string) : M (option ((mword 3 * mword 6))) :=
+ let _s1808_ := _s1807_ in
+ (if string_startswith _s1808_ "c.srli" then
+ (match (string_drop _s1808_ (projT1 (string_length "c.srli"))) with
+ | _s1809_ =>
+ (spc_matches_prefix _s1809_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1291_ _)) =>
- (match (string_drop _s1290_ _s1291_) with
- | _s1292_ =>
- (creg_name_matches_prefix _s1292_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1810_ _) =>
+ (match (string_drop _s1809_ _s1810_) with
+ | _s1811_ =>
+ (creg_name_matches_prefix _s1811_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1293_ _)) =>
- (match (string_drop _s1292_ _s1293_) with
- | _s1294_ =>
- (sep_matches_prefix _s1294_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1812_ _) =>
+ (match (string_drop _s1811_ _s1812_) with
+ | _s1813_ =>
+ (sep_matches_prefix _s1813_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1295_ _)) =>
- match (string_drop _s1294_ _s1295_) with
- | _s1296_ =>
- match (hex_bits_6_matches_prefix _s1296_) with
- | Some ((shamt, existT _ _s1297_ _)) =>
- let p0_ := string_drop _s1296_ _s1297_ in
- if ((generic_eq p0_ "")) then Some ((rsd, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1814_ _) =>
+ match (string_drop _s1813_ _s1814_) with
+ | _s1815_ =>
+ match (hex_bits_6_matches_prefix _s1815_) with
+ | Some (shamt, existT _ _s1816_ _) =>
+ let p0_ := string_drop _s1815_ _s1816_ in
+ if generic_eq p0_ "" then Some (rsd, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- else returnm (None : option ((mword 3 * mword 6))))
+ else returnm None)
: M (option ((mword 3 * mword 6))).
-Definition _s1275_ (_s1276_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1277_ := _s1276_ in
- (if ((string_startswith _s1277_ "c.lui")) then
- (match (string_drop _s1277_ (projT1 (string_length "c.lui"))) with
- | _s1278_ =>
- (spc_matches_prefix _s1278_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1794_ (_s1795_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1796_ := _s1795_ in
+ (if string_startswith _s1796_ "c.lui" then
+ (match (string_drop _s1796_ (projT1 (string_length "c.lui"))) with
+ | _s1797_ =>
+ (spc_matches_prefix _s1797_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1279_ _)) =>
- (match (string_drop _s1278_ _s1279_) with
- | _s1280_ =>
- (reg_name_matches_prefix _s1280_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1798_ _) =>
+ (match (string_drop _s1797_ _s1798_) with
+ | _s1799_ =>
+ (reg_name_matches_prefix _s1799_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1281_ _)) =>
- (match (string_drop _s1280_ _s1281_) with
- | _s1282_ =>
- (sep_matches_prefix _s1282_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1800_ _) =>
+ (match (string_drop _s1799_ _s1800_) with
+ | _s1801_ =>
+ (sep_matches_prefix _s1801_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1283_ _)) =>
- match (string_drop _s1282_ _s1283_) with
- | _s1284_ =>
- match (hex_bits_6_matches_prefix _s1284_) with
- | Some ((imm, existT _ _s1285_ _)) =>
- let p0_ := string_drop _s1284_ _s1285_ in
- if ((generic_eq p0_ "")) then Some ((rd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1802_ _) =>
+ match (string_drop _s1801_ _s1802_) with
+ | _s1803_ =>
+ match (hex_bits_6_matches_prefix _s1803_) with
+ | Some (imm, existT _ _s1804_ _) =>
+ let p0_ := string_drop _s1803_ _s1804_ in
+ if generic_eq p0_ "" then Some (rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1267_ (_s1268_ : string)
-: M (option (mword 6)) :=
-
- let _s1269_ := _s1268_ in
- (if ((string_startswith _s1269_ "c.addi16sp")) then
- (match (string_drop _s1269_ (projT1 (string_length "c.addi16sp"))) with
- | _s1270_ =>
- (spc_matches_prefix _s1270_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s1271_ _)) =>
- match (string_drop _s1270_ _s1271_) with
- | _s1272_ =>
- match (hex_bits_6_matches_prefix _s1272_) with
- | Some ((imm, existT _ _s1273_ _)) =>
- let p0_ := string_drop _s1272_ _s1273_ in
- if ((generic_eq p0_ "")) then Some (imm)
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 6))
+Definition _s1786_ (_s1787_ : string) : M (option (mword 6)) :=
+ let _s1788_ := _s1787_ in
+ (if string_startswith _s1788_ "c.addi16sp" then
+ (match (string_drop _s1788_ (projT1 (string_length "c.addi16sp"))) with
+ | _s1789_ =>
+ (spc_matches_prefix _s1789_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s1790_ _) =>
+ match (string_drop _s1789_ _s1790_) with
+ | _s1791_ =>
+ match (hex_bits_6_matches_prefix _s1791_) with
+ | Some (imm, existT _ _s1792_ _) =>
+ let p0_ := string_drop _s1791_ _s1792_ in
+ if generic_eq p0_ "" then Some imm
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 6))
- else returnm (None : option (mword 6)))
+ else returnm None)
: M (option (mword 6)).
-Definition _s1255_ (_s1256_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1257_ := _s1256_ in
- (if ((string_startswith _s1257_ "c.li")) then
- (match (string_drop _s1257_ (projT1 (string_length "c.li"))) with
- | _s1258_ =>
- (spc_matches_prefix _s1258_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1774_ (_s1775_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1776_ := _s1775_ in
+ (if string_startswith _s1776_ "c.li" then
+ (match (string_drop _s1776_ (projT1 (string_length "c.li"))) with
+ | _s1777_ =>
+ (spc_matches_prefix _s1777_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1259_ _)) =>
- (match (string_drop _s1258_ _s1259_) with
- | _s1260_ =>
- (reg_name_matches_prefix _s1260_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1778_ _) =>
+ (match (string_drop _s1777_ _s1778_) with
+ | _s1779_ =>
+ (reg_name_matches_prefix _s1779_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1261_ _)) =>
- (match (string_drop _s1260_ _s1261_) with
- | _s1262_ =>
- (sep_matches_prefix _s1262_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1780_ _) =>
+ (match (string_drop _s1779_ _s1780_) with
+ | _s1781_ =>
+ (sep_matches_prefix _s1781_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1263_ _)) =>
- match (string_drop _s1262_ _s1263_) with
- | _s1264_ =>
- match (hex_bits_6_matches_prefix _s1264_) with
- | Some ((imm, existT _ _s1265_ _)) =>
- let p0_ := string_drop _s1264_ _s1265_ in
- if ((generic_eq p0_ "")) then Some ((rd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1782_ _) =>
+ match (string_drop _s1781_ _s1782_) with
+ | _s1783_ =>
+ match (hex_bits_6_matches_prefix _s1783_) with
+ | Some (imm, existT _ _s1784_ _) =>
+ let p0_ := string_drop _s1783_ _s1784_ in
+ if generic_eq p0_ "" then Some (rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1243_ (_s1244_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1245_ := _s1244_ in
- (if ((string_startswith _s1245_ "c.addiw")) then
- (match (string_drop _s1245_ (projT1 (string_length "c.addiw"))) with
- | _s1246_ =>
- (spc_matches_prefix _s1246_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1762_ (_s1763_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1764_ := _s1763_ in
+ (if string_startswith _s1764_ "c.addiw" then
+ (match (string_drop _s1764_ (projT1 (string_length "c.addiw"))) with
+ | _s1765_ =>
+ (spc_matches_prefix _s1765_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1247_ _)) =>
- (match (string_drop _s1246_ _s1247_) with
- | _s1248_ =>
- (reg_name_matches_prefix _s1248_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1766_ _) =>
+ (match (string_drop _s1765_ _s1766_) with
+ | _s1767_ =>
+ (reg_name_matches_prefix _s1767_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1249_ _)) =>
- (match (string_drop _s1248_ _s1249_) with
- | _s1250_ =>
- (sep_matches_prefix _s1250_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1768_ _) =>
+ (match (string_drop _s1767_ _s1768_) with
+ | _s1769_ =>
+ (sep_matches_prefix _s1769_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1251_ _)) =>
- match (string_drop _s1250_ _s1251_) with
- | _s1252_ =>
- match (hex_bits_6_matches_prefix _s1252_) with
- | Some ((imm, existT _ _s1253_ _)) =>
- let p0_ := string_drop _s1252_ _s1253_ in
- if ((generic_eq p0_ "")) then Some ((rsd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1770_ _) =>
+ match (string_drop _s1769_ _s1770_) with
+ | _s1771_ =>
+ match (hex_bits_6_matches_prefix _s1771_) with
+ | Some (imm, existT _ _s1772_ _) =>
+ let p0_ := string_drop _s1771_ _s1772_ in
+ if generic_eq p0_ "" then Some (rsd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1235_ (_s1236_ : string)
-: M (option (mword 11)) :=
-
- let _s1237_ := _s1236_ in
- (if ((string_startswith _s1237_ "c.jal")) then
- (match (string_drop _s1237_ (projT1 (string_length "c.jal"))) with
- | _s1238_ =>
- (spc_matches_prefix _s1238_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s1239_ _)) =>
- match (string_drop _s1238_ _s1239_) with
- | _s1240_ =>
- match (hex_bits_12_matches_prefix _s1240_) with
- | Some ((v__802, existT _ _s1241_ _)) =>
- if ((eq_vec (subrange_vec_dec v__802 0 0)
- (vec_of_bits [B0] : mword (0 - 0 + 1)))) then
- let imm : mword 11 := subrange_vec_dec v__802 11 1 in
- let imm : mword 11 := subrange_vec_dec v__802 11 1 in
- let p0_ := string_drop _s1240_ _s1241_ in
- if ((generic_eq p0_ "")) then Some (imm)
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 11))
+Definition _s1754_ (_s1755_ : string) : M (option (mword 11)) :=
+ let _s1756_ := _s1755_ in
+ (if string_startswith _s1756_ "c.jal" then
+ (match (string_drop _s1756_ (projT1 (string_length "c.jal"))) with
+ | _s1757_ =>
+ (spc_matches_prefix _s1757_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s1758_ _) =>
+ match (string_drop _s1757_ _s1758_) with
+ | _s1759_ =>
+ match (hex_bits_12_matches_prefix _s1759_) with
+ | Some (v__1346, existT _ _s1760_ _) =>
+ if eq_vec (subrange_vec_dec v__1346 0 0) ('b"0" : mword (0 - 0 + 1))
+ then
+ let imm : mword 11 := subrange_vec_dec v__1346 11 1 in
+ let imm : mword 11 := subrange_vec_dec v__1346 11 1 in
+ let p0_ := string_drop _s1759_ _s1760_ in
+ if generic_eq p0_ "" then Some imm
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 11))
- else returnm (None : option (mword 11)))
+ else returnm None)
: M (option (mword 11)).
-Definition _s1223_ (_s1224_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s1225_ := _s1224_ in
- (if ((string_startswith _s1225_ "c.addi")) then
- (match (string_drop _s1225_ (projT1 (string_length "c.addi"))) with
- | _s1226_ =>
- (spc_matches_prefix _s1226_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1742_ (_s1743_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s1744_ := _s1743_ in
+ (if string_startswith _s1744_ "c.addi" then
+ (match (string_drop _s1744_ (projT1 (string_length "c.addi"))) with
+ | _s1745_ =>
+ (spc_matches_prefix _s1745_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1227_ _)) =>
- (match (string_drop _s1226_ _s1227_) with
- | _s1228_ =>
- (reg_name_matches_prefix _s1228_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1746_ _) =>
+ (match (string_drop _s1745_ _s1746_) with
+ | _s1747_ =>
+ (reg_name_matches_prefix _s1747_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s1229_ _)) =>
- (match (string_drop _s1228_ _s1229_) with
- | _s1230_ =>
- (sep_matches_prefix _s1230_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s1748_ _) =>
+ (match (string_drop _s1747_ _s1748_) with
+ | _s1749_ =>
+ (sep_matches_prefix _s1749_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1231_ _)) =>
- match (string_drop _s1230_ _s1231_) with
- | _s1232_ =>
- match (hex_bits_6_matches_prefix _s1232_) with
- | Some ((nzi, existT _ _s1233_ _)) =>
- let p0_ := string_drop _s1232_ _s1233_ in
- if ((generic_eq p0_ "")) then Some ((rsd, nzi))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1750_ _) =>
+ match (string_drop _s1749_ _s1750_) with
+ | _s1751_ =>
+ match (hex_bits_6_matches_prefix _s1751_) with
+ | Some (nzi, existT _ _s1752_ _) =>
+ let p0_ := string_drop _s1751_ _s1752_ in
+ if generic_eq p0_ "" then Some (rsd, nzi)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s1207_ (_s1208_ : string)
-: M (option ((mword 3 * mword 3 * mword 5))) :=
-
- let _s1209_ := _s1208_ in
- (if ((string_startswith _s1209_ "c.sd")) then
- (match (string_drop _s1209_ (projT1 (string_length "c.sd"))) with
- | _s1210_ =>
- (spc_matches_prefix _s1210_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1726_ (_s1727_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s1728_ := _s1727_ in
+ (if string_startswith _s1728_ "c.sd" then
+ (match (string_drop _s1728_ (projT1 (string_length "c.sd"))) with
+ | _s1729_ =>
+ (spc_matches_prefix _s1729_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1211_ _)) =>
- (match (string_drop _s1210_ _s1211_) with
- | _s1212_ =>
- (creg_name_matches_prefix _s1212_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1730_ _) =>
+ (match (string_drop _s1729_ _s1730_) with
+ | _s1731_ =>
+ (creg_name_matches_prefix _s1731_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsc1, existT _ _s1213_ _)) =>
- (match (string_drop _s1212_ _s1213_) with
- | _s1214_ =>
- (sep_matches_prefix _s1214_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc1, existT _ _s1732_ _) =>
+ (match (string_drop _s1731_ _s1732_) with
+ | _s1733_ =>
+ (sep_matches_prefix _s1733_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1215_ _)) =>
- (match (string_drop _s1214_ _s1215_) with
- | _s1216_ =>
- (creg_name_matches_prefix _s1216_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1734_ _) =>
+ (match (string_drop _s1733_ _s1734_) with
+ | _s1735_ =>
+ (creg_name_matches_prefix _s1735_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc2, existT _ _s1217_ _)) =>
- (match (string_drop _s1216_ _s1217_) with
- | _s1218_ =>
- (sep_matches_prefix _s1218_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc2, existT _ _s1736_ _) =>
+ (match (string_drop _s1735_ _s1736_) with
+ | _s1737_ =>
+ (sep_matches_prefix _s1737_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s1219_ _)) =>
- match (string_drop _s1218_ _s1219_) with
- | _s1220_ =>
- match (hex_bits_8_matches_prefix _s1220_) with
- | Some ((v__804, existT _ _s1221_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__804 2 0)
- (vec_of_bits [B0;B0;B0]
- : mword (2 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__804 7 3 in
- let uimm : mword 5 :=
- subrange_vec_dec v__804 7 3 in
- let p0_ :=
- string_drop _s1220_ _s1221_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rsc1, rsc2, uimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s1738_ _) =>
+ match (string_drop _s1737_ _s1738_) with
+ | _s1739_ =>
+ match (hex_bits_8_matches_prefix _s1739_) with
+ | Some (v__1348, existT _ _s1740_ _) =>
+ if eq_vec (subrange_vec_dec v__1348 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1348 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1348 7 3 in
+ let p0_ := string_drop _s1739_ _s1740_ in
+ if generic_eq p0_ "" then
+ Some (rsc1, rsc2, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5))).
-Definition _s1191_ (_s1192_ : string)
-: M (option ((mword 3 * mword 3 * mword 5))) :=
-
- let _s1193_ := _s1192_ in
- (if ((string_startswith _s1193_ "c.sw")) then
- (match (string_drop _s1193_ (projT1 (string_length "c.sw"))) with
- | _s1194_ =>
- (spc_matches_prefix _s1194_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1710_ (_s1711_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s1712_ := _s1711_ in
+ (if string_startswith _s1712_ "c.sw" then
+ (match (string_drop _s1712_ (projT1 (string_length "c.sw"))) with
+ | _s1713_ =>
+ (spc_matches_prefix _s1713_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1195_ _)) =>
- (match (string_drop _s1194_ _s1195_) with
- | _s1196_ =>
- (creg_name_matches_prefix _s1196_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1714_ _) =>
+ (match (string_drop _s1713_ _s1714_) with
+ | _s1715_ =>
+ (creg_name_matches_prefix _s1715_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsc1, existT _ _s1197_ _)) =>
- (match (string_drop _s1196_ _s1197_) with
- | _s1198_ =>
- (sep_matches_prefix _s1198_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc1, existT _ _s1716_ _) =>
+ (match (string_drop _s1715_ _s1716_) with
+ | _s1717_ =>
+ (sep_matches_prefix _s1717_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1199_ _)) =>
- (match (string_drop _s1198_ _s1199_) with
- | _s1200_ =>
- (creg_name_matches_prefix _s1200_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1718_ _) =>
+ (match (string_drop _s1717_ _s1718_) with
+ | _s1719_ =>
+ (creg_name_matches_prefix _s1719_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc2, existT _ _s1201_ _)) =>
- (match (string_drop _s1200_ _s1201_) with
- | _s1202_ =>
- (sep_matches_prefix _s1202_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc2, existT _ _s1720_ _) =>
+ (match (string_drop _s1719_ _s1720_) with
+ | _s1721_ =>
+ (sep_matches_prefix _s1721_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s1203_ _)) =>
- match (string_drop _s1202_ _s1203_) with
- | _s1204_ =>
- match (hex_bits_7_matches_prefix _s1204_) with
- | Some ((v__806, existT _ _s1205_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__806 1 0)
- (vec_of_bits [B0;B0]
- : mword (1 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__806 6 2 in
- let uimm : mword 5 :=
- subrange_vec_dec v__806 6 2 in
- let p0_ :=
- string_drop _s1204_ _s1205_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rsc1, rsc2, uimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s1722_ _) =>
+ match (string_drop _s1721_ _s1722_) with
+ | _s1723_ =>
+ match (hex_bits_7_matches_prefix _s1723_) with
+ | Some (v__1350, existT _ _s1724_ _) =>
+ if eq_vec (subrange_vec_dec v__1350 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1350 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1350 6 2 in
+ let p0_ := string_drop _s1723_ _s1724_ in
+ if generic_eq p0_ "" then
+ Some (rsc1, rsc2, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5))).
-Definition _s1175_ (_s1176_ : string)
-: M (option ((mword 3 * mword 3 * mword 5))) :=
-
- let _s1177_ := _s1176_ in
- (if ((string_startswith _s1177_ "c.ld")) then
- (match (string_drop _s1177_ (projT1 (string_length "c.ld"))) with
- | _s1178_ =>
- (spc_matches_prefix _s1178_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1694_ (_s1695_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s1696_ := _s1695_ in
+ (if string_startswith _s1696_ "c.ld" then
+ (match (string_drop _s1696_ (projT1 (string_length "c.ld"))) with
+ | _s1697_ =>
+ (spc_matches_prefix _s1697_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1179_ _)) =>
- (match (string_drop _s1178_ _s1179_) with
- | _s1180_ =>
- (creg_name_matches_prefix _s1180_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1698_ _) =>
+ (match (string_drop _s1697_ _s1698_) with
+ | _s1699_ =>
+ (creg_name_matches_prefix _s1699_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s1181_ _)) =>
- (match (string_drop _s1180_ _s1181_) with
- | _s1182_ =>
- (sep_matches_prefix _s1182_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s1700_ _) =>
+ (match (string_drop _s1699_ _s1700_) with
+ | _s1701_ =>
+ (sep_matches_prefix _s1701_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1183_ _)) =>
- (match (string_drop _s1182_ _s1183_) with
- | _s1184_ =>
- (creg_name_matches_prefix _s1184_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1702_ _) =>
+ (match (string_drop _s1701_ _s1702_) with
+ | _s1703_ =>
+ (creg_name_matches_prefix _s1703_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc, existT _ _s1185_ _)) =>
- (match (string_drop _s1184_ _s1185_) with
- | _s1186_ =>
- (sep_matches_prefix _s1186_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc, existT _ _s1704_ _) =>
+ (match (string_drop _s1703_ _s1704_) with
+ | _s1705_ =>
+ (sep_matches_prefix _s1705_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s1187_ _)) =>
- match (string_drop _s1186_ _s1187_) with
- | _s1188_ =>
- match (hex_bits_8_matches_prefix _s1188_) with
- | Some ((v__808, existT _ _s1189_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__808 2 0)
- (vec_of_bits [B0;B0;B0]
- : mword (2 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__808 7 3 in
- let uimm : mword 5 :=
- subrange_vec_dec v__808 7 3 in
- let p0_ :=
- string_drop _s1188_ _s1189_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rdc, rsc, uimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s1706_ _) =>
+ match (string_drop _s1705_ _s1706_) with
+ | _s1707_ =>
+ match (hex_bits_8_matches_prefix _s1707_) with
+ | Some (v__1352, existT _ _s1708_ _) =>
+ if eq_vec (subrange_vec_dec v__1352 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1352 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1352 7 3 in
+ let p0_ := string_drop _s1707_ _s1708_ in
+ if generic_eq p0_ "" then
+ Some (rdc, rsc, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5))).
-Definition _s1159_ (_s1160_ : string)
-: M (option ((mword 3 * mword 3 * mword 5))) :=
-
- let _s1161_ := _s1160_ in
- (if ((string_startswith _s1161_ "c.lw")) then
- (match (string_drop _s1161_ (projT1 (string_length "c.lw"))) with
- | _s1162_ =>
- (spc_matches_prefix _s1162_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1678_ (_s1679_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s1680_ := _s1679_ in
+ (if string_startswith _s1680_ "c.lw" then
+ (match (string_drop _s1680_ (projT1 (string_length "c.lw"))) with
+ | _s1681_ =>
+ (spc_matches_prefix _s1681_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1163_ _)) =>
- (match (string_drop _s1162_ _s1163_) with
- | _s1164_ =>
- (creg_name_matches_prefix _s1164_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1682_ _) =>
+ (match (string_drop _s1681_ _s1682_) with
+ | _s1683_ =>
+ (creg_name_matches_prefix _s1683_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s1165_ _)) =>
- (match (string_drop _s1164_ _s1165_) with
- | _s1166_ =>
- (sep_matches_prefix _s1166_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s1684_ _) =>
+ (match (string_drop _s1683_ _s1684_) with
+ | _s1685_ =>
+ (sep_matches_prefix _s1685_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1167_ _)) =>
- (match (string_drop _s1166_ _s1167_) with
- | _s1168_ =>
- (creg_name_matches_prefix _s1168_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1686_ _) =>
+ (match (string_drop _s1685_ _s1686_) with
+ | _s1687_ =>
+ (creg_name_matches_prefix _s1687_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc, existT _ _s1169_ _)) =>
- (match (string_drop _s1168_ _s1169_) with
- | _s1170_ =>
- (sep_matches_prefix _s1170_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc, existT _ _s1688_ _) =>
+ (match (string_drop _s1687_ _s1688_) with
+ | _s1689_ =>
+ (sep_matches_prefix _s1689_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s1171_ _)) =>
- match (string_drop _s1170_ _s1171_) with
- | _s1172_ =>
- match (hex_bits_7_matches_prefix _s1172_) with
- | Some ((v__810, existT _ _s1173_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__810 1 0)
- (vec_of_bits [B0;B0]
- : mword (1 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__810 6 2 in
- let uimm : mword 5 :=
- subrange_vec_dec v__810 6 2 in
- let p0_ :=
- string_drop _s1172_ _s1173_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rdc, rsc, uimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s1690_ _) =>
+ match (string_drop _s1689_ _s1690_) with
+ | _s1691_ =>
+ match (hex_bits_7_matches_prefix _s1691_) with
+ | Some (v__1354, existT _ _s1692_ _) =>
+ if eq_vec (subrange_vec_dec v__1354 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1354 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1354 6 2 in
+ let p0_ := string_drop _s1691_ _s1692_ in
+ if generic_eq p0_ "" then
+ Some (rdc, rsc, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5))).
-Definition _s1147_ (_s1148_ : string)
-: M (option ((mword 3 * mword 8))) :=
-
- let _s1149_ := _s1148_ in
- (if ((string_startswith _s1149_ "c.addi4spn")) then
- (match (string_drop _s1149_ (projT1 (string_length "c.addi4spn"))) with
- | _s1150_ =>
- (spc_matches_prefix _s1150_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1666_ (_s1667_ : string) : M (option ((mword 3 * mword 8))) :=
+ let _s1668_ := _s1667_ in
+ (if string_startswith _s1668_ "c.addi4spn" then
+ (match (string_drop _s1668_ (projT1 (string_length "c.addi4spn"))) with
+ | _s1669_ =>
+ (spc_matches_prefix _s1669_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1151_ _)) =>
- (match (string_drop _s1150_ _s1151_) with
- | _s1152_ =>
- (creg_name_matches_prefix _s1152_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1670_ _) =>
+ (match (string_drop _s1669_ _s1670_) with
+ | _s1671_ =>
+ (creg_name_matches_prefix _s1671_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s1153_ _)) =>
- (match (string_drop _s1152_ _s1153_) with
- | _s1154_ =>
- (sep_matches_prefix _s1154_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s1672_ _) =>
+ (match (string_drop _s1671_ _s1672_) with
+ | _s1673_ =>
+ (sep_matches_prefix _s1673_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1155_ _)) =>
- match (string_drop _s1154_ _s1155_) with
- | _s1156_ =>
- match (hex_bits_10_matches_prefix _s1156_) with
- | Some ((v__812, existT _ _s1157_ _)) =>
- if ((eq_vec (subrange_vec_dec v__812 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1))))
- then
- let nzimm : mword 8 := subrange_vec_dec v__812 9 2 in
- let nzimm : mword 8 := subrange_vec_dec v__812 9 2 in
- let p0_ := string_drop _s1156_ _s1157_ in
- if ((generic_eq p0_ "")) then Some ((rdc, nzimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1674_ _) =>
+ match (string_drop _s1673_ _s1674_) with
+ | _s1675_ =>
+ match (hex_bits_10_matches_prefix _s1675_) with
+ | Some (v__1356, existT _ _s1676_ _) =>
+ if eq_vec (subrange_vec_dec v__1356 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let nzimm : mword 8 := subrange_vec_dec v__1356 9 2 in
+ let nzimm : mword 8 := subrange_vec_dec v__1356 9 2 in
+ let p0_ := string_drop _s1675_ _s1676_ in
+ if generic_eq p0_ "" then Some (rdc, nzimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- else returnm (None : option ((mword 3 * mword 8))))
+ else returnm None)
: M (option ((mword 3 * mword 8))).
-Definition _s1123_ (_s1124_ : string)
+Definition _s1640_ (_s1641_ : string)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1124_ with
- | _s1125_ =>
- (amo_mnemonic_matches_prefix _s1125_) >>= fun w__0 : option ((amoop * {n : Z & ArithFact (n >=
+ (match _s1641_ with
+ | _s1642_ =>
+ (amo_mnemonic_matches_prefix _s1642_) >>= fun w__0 : option ((amoop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1126_ _)) =>
- let _s1127_ := string_drop _s1125_ _s1126_ in
- (if ((string_startswith _s1127_ ".")) then
- (match (string_drop _s1127_ (projT1 (string_length "."))) with
- | _s1128_ =>
- (size_mnemonic_matches_prefix _s1128_) >>= fun w__1 : option ((word_width * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s1643_ _) =>
+ let _s1644_ := string_drop _s1642_ _s1643_ in
+ (if string_startswith _s1644_ "." then
+ (match (string_drop _s1644_ (projT1 (string_length "."))) with
+ | _s1645_ =>
+ (size_mnemonic_matches_prefix _s1645_) >>= fun w__1 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((width, existT _ _s1129_ _)) =>
- (match (string_drop _s1128_ _s1129_) with
- | _s1130_ =>
- (maybe_aq_matches_prefix _s1130_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (width, existT _ _s1646_ _) =>
+ (match (string_drop _s1645_ _s1646_) with
+ | _s1647_ =>
+ (maybe_aq_matches_prefix _s1647_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((aq, existT _ _s1131_ _)) =>
- (match (string_drop _s1130_ _s1131_) with
- | _s1132_ =>
- (maybe_rl_matches_prefix _s1132_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s1648_ _) =>
+ (match (string_drop _s1647_ _s1648_) with
+ | _s1649_ =>
+ (maybe_rl_matches_prefix _s1649_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rl, existT _ _s1133_ _)) =>
- (match (string_drop _s1132_ _s1133_) with
- | _s1134_ =>
- (spc_matches_prefix _s1134_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s1650_ _) =>
+ (match (string_drop _s1649_ _s1650_) with
+ | _s1651_ =>
+ (spc_matches_prefix _s1651_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s1135_ _)) =>
- (match (string_drop _s1134_ _s1135_) with
- | _s1136_ =>
- (reg_name_matches_prefix _s1136_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1652_ _) =>
+ (match (string_drop _s1651_ _s1652_) with
+ | _s1653_ =>
+ (reg_name_matches_prefix _s1653_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((rd, existT _ _s1137_ _)) =>
- (match (string_drop _s1136_ _s1137_) with
- | _s1138_ =>
- (sep_matches_prefix _s1138_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1654_ _) =>
+ (match (string_drop _s1653_ _s1654_) with
+ | _s1655_ =>
+ (sep_matches_prefix _s1655_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s1139_ _)) =>
- (match (string_drop _s1138_ _s1139_) with
- | _s1140_ =>
- (reg_name_matches_prefix _s1140_) >>= fun w__7 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1656_ _) =>
+ (match (string_drop _s1655_ _s1656_) with
+ | _s1657_ =>
+ (reg_name_matches_prefix _s1657_) >>= fun w__7 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
- | Some
- ((rs1, existT _ _s1141_ _)) =>
- (match (string_drop _s1140_
- _s1141_) with
- | _s1142_ =>
+ | Some (rs2, existT _ _s1658_ _) =>
+ (match (string_drop _s1657_
+ _s1658_) with
+ | _s1659_ =>
(sep_matches_prefix
- _s1142_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=
+ _s1659_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((tt, existT _ _s1143_ _)) =>
- (match (string_drop
- _s1142_
- _s1143_) with
- | _s1144_ =>
- (reg_name_matches_prefix
- _s1144_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=
- 0)})) =>
- returnm ((match w__9 with
- | Some
- ((rs2, existT _ _s1145_ _)) =>
- let p0_ :=
- string_drop
- _s1144_
- _s1145_ in
- if ((generic_eq
- p0_
- ""))
- then
- Some
- ((op, width, aq, rl, rd, rs1, rs2))
- else
+ (tt, existT _ _s1660_ _) =>
+ let _s1661_ :=
+ string_drop _s1659_
+ _s1660_ in
+ (if string_startswith
+ _s1661_ "(" then
+ (match (string_drop
+ _s1661_
+ (projT1
+ (string_length
+ "("))) with
+ | _s1662_ =>
+ (reg_name_matches_prefix
+ _s1662_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__9 with
+ | Some
+ (rs1, existT _ _s1663_ _) =>
+ let _s1664_ :=
+ string_drop
+ _s1662_
+ _s1663_ in
+ if string_startswith
+ _s1664_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s1664_
+ (projT1
+ (string_length
+ ")")) in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (op, width, aq, rl, rd, rs2, rs1)
+ else
+ None
+ else
+ None
+ | _ =>
None
- | _ =>
- None
- end)
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- end)
+ end)
+ end)
+ : M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ else returnm None)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- else
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5))).
-Definition _s1101_ (_s1102_ : string)
+Definition _s1618_ (_s1619_ : string)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s1103_ := _s1102_ in
- (if ((string_startswith _s1103_ "sc.")) then
- (match (string_drop _s1103_ (projT1 (string_length "sc."))) with
- | _s1104_ =>
- (size_mnemonic_matches_prefix _s1104_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s1620_ := _s1619_ in
+ (if string_startswith _s1620_ "sc." then
+ (match (string_drop _s1620_ (projT1 (string_length "sc."))) with
+ | _s1621_ =>
+ (size_mnemonic_matches_prefix _s1621_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s1105_ _)) =>
- (match (string_drop _s1104_ _s1105_) with
- | _s1106_ =>
- (maybe_aq_matches_prefix _s1106_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s1622_ _) =>
+ (match (string_drop _s1621_ _s1622_) with
+ | _s1623_ =>
+ (maybe_aq_matches_prefix _s1623_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s1107_ _)) =>
- (match (string_drop _s1106_ _s1107_) with
- | _s1108_ =>
- (maybe_rl_matches_prefix _s1108_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s1624_ _) =>
+ (match (string_drop _s1623_ _s1624_) with
+ | _s1625_ =>
+ (maybe_rl_matches_prefix _s1625_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s1109_ _)) =>
- (match (string_drop _s1108_ _s1109_) with
- | _s1110_ =>
- (spc_matches_prefix _s1110_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s1626_ _) =>
+ (match (string_drop _s1625_ _s1626_) with
+ | _s1627_ =>
+ (spc_matches_prefix _s1627_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1111_ _)) =>
- (match (string_drop _s1110_ _s1111_) with
- | _s1112_ =>
- (reg_name_matches_prefix _s1112_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1628_ _) =>
+ (match (string_drop _s1627_ _s1628_) with
+ | _s1629_ =>
+ (reg_name_matches_prefix _s1629_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rd, existT _ _s1113_ _)) =>
- (match (string_drop _s1112_ _s1113_) with
- | _s1114_ =>
- (sep_matches_prefix _s1114_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1630_ _) =>
+ (match (string_drop _s1629_ _s1630_) with
+ | _s1631_ =>
+ (sep_matches_prefix _s1631_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1115_ _)) =>
- (match (string_drop _s1114_ _s1115_) with
- | _s1116_ =>
- (reg_name_matches_prefix _s1116_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1632_ _) =>
+ (match (string_drop _s1631_ _s1632_) with
+ | _s1633_ =>
+ (reg_name_matches_prefix _s1633_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((rs1, existT _ _s1117_ _)) =>
- (match (string_drop _s1116_ _s1117_) with
- | _s1118_ =>
- (sep_matches_prefix _s1118_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1634_ _) =>
+ (match (string_drop _s1633_ _s1634_) with
+ | _s1635_ =>
+ (sep_matches_prefix _s1635_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
- | Some ((tt, existT _ _s1119_ _)) =>
- (match (string_drop _s1118_
- _s1119_) with
- | _s1120_ =>
+ | Some (tt, existT _ _s1636_ _) =>
+ (match (string_drop _s1635_
+ _s1636_) with
+ | _s1637_ =>
(reg_name_matches_prefix
- _s1120_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s1637_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__8 with
- | Some
- ((rs2, existT _ _s1121_ _)) =>
- let p0_ :=
- string_drop
- _s1120_
- _s1121_ in
- if ((generic_eq
- p0_ ""))
- then
- Some
- ((size, aq, rl, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__8 with
+ | Some
+ (rs2, existT _ _s1638_ _) =>
+ let p0_ :=
+ string_drop
+ _s1637_
+ _s1638_ in
+ if generic_eq
+ p0_ ""
+ then
+ Some
+ (size, aq, rl, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5))).
-Definition _s1083_ (_s1084_ : string)
-: M (option ((word_width * bool * bool * mword 5 * mword 5))) :=
-
- let _s1085_ := _s1084_ in
- (if ((string_startswith _s1085_ "lr.")) then
- (match (string_drop _s1085_ (projT1 (string_length "lr."))) with
- | _s1086_ =>
- (size_mnemonic_matches_prefix _s1086_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+Definition _s1600_ (_s1601_ : string) : M (option ((word_width * bool * bool * mword 5 * mword 5))) :=
+ let _s1602_ := _s1601_ in
+ (if string_startswith _s1602_ "lr." then
+ (match (string_drop _s1602_ (projT1 (string_length "lr."))) with
+ | _s1603_ =>
+ (size_mnemonic_matches_prefix _s1603_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s1087_ _)) =>
- (match (string_drop _s1086_ _s1087_) with
- | _s1088_ =>
- (maybe_aq_matches_prefix _s1088_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s1604_ _) =>
+ (match (string_drop _s1603_ _s1604_) with
+ | _s1605_ =>
+ (maybe_aq_matches_prefix _s1605_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s1089_ _)) =>
- (match (string_drop _s1088_ _s1089_) with
- | _s1090_ =>
- (maybe_rl_matches_prefix _s1090_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s1606_ _) =>
+ (match (string_drop _s1605_ _s1606_) with
+ | _s1607_ =>
+ (maybe_rl_matches_prefix _s1607_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s1091_ _)) =>
- (match (string_drop _s1090_ _s1091_) with
- | _s1092_ =>
- (spc_matches_prefix _s1092_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s1608_ _) =>
+ (match (string_drop _s1607_ _s1608_) with
+ | _s1609_ =>
+ (spc_matches_prefix _s1609_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1093_ _)) =>
- (match (string_drop _s1092_ _s1093_) with
- | _s1094_ =>
- (reg_name_matches_prefix _s1094_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1610_ _) =>
+ (match (string_drop _s1609_ _s1610_) with
+ | _s1611_ =>
+ (reg_name_matches_prefix _s1611_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rd, existT _ _s1095_ _)) =>
- (match (string_drop _s1094_ _s1095_) with
- | _s1096_ =>
- (sep_matches_prefix _s1096_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1612_ _) =>
+ (match (string_drop _s1611_ _s1612_) with
+ | _s1613_ =>
+ (sep_matches_prefix _s1613_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1097_ _)) =>
- (match (string_drop _s1096_ _s1097_) with
- | _s1098_ =>
- (reg_name_matches_prefix _s1098_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1614_ _) =>
+ (match (string_drop _s1613_ _s1614_) with
+ | _s1615_ =>
+ (reg_name_matches_prefix _s1615_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs1, existT _ _s1099_ _)) =>
- let p0_ :=
- string_drop _s1098_ _s1099_ in
- if ((generic_eq p0_ "")) then
- Some
- ((size, aq, rl, rd, rs1))
- else None
- | _ => None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs1, existT _ _s1616_ _) =>
+ let p0_ :=
+ string_drop _s1615_ _s1616_ in
+ if generic_eq p0_ "" then
+ Some (size, aq, rl, rd, rs1)
+ else None
+ | _ => None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ => returnm (None : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ => returnm (None : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- else returnm (None : option ((word_width * bool * bool * mword 5 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 5))).
-Definition _s1071_ (_s1072_ : string)
-: M (option ((mword 5 * mword 5))) :=
-
- let _s1073_ := _s1072_ in
- (if ((string_startswith _s1073_ "sfence.vma")) then
- (match (string_drop _s1073_ (projT1 (string_length "sfence.vma"))) with
- | _s1074_ =>
- (spc_matches_prefix _s1074_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1588_ (_s1589_ : string) : M (option ((mword 5 * mword 5))) :=
+ let _s1590_ := _s1589_ in
+ (if string_startswith _s1590_ "sfence.vma" then
+ (match (string_drop _s1590_ (projT1 (string_length "sfence.vma"))) with
+ | _s1591_ =>
+ (spc_matches_prefix _s1591_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1075_ _)) =>
- (match (string_drop _s1074_ _s1075_) with
- | _s1076_ =>
- (reg_name_matches_prefix _s1076_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1592_ _) =>
+ (match (string_drop _s1591_ _s1592_) with
+ | _s1593_ =>
+ (reg_name_matches_prefix _s1593_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs1, existT _ _s1077_ _)) =>
- (match (string_drop _s1076_ _s1077_) with
- | _s1078_ =>
- (sep_matches_prefix _s1078_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1594_ _) =>
+ (match (string_drop _s1593_ _s1594_) with
+ | _s1595_ =>
+ (sep_matches_prefix _s1595_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1079_ _)) =>
- (match (string_drop _s1078_ _s1079_) with
- | _s1080_ =>
- (reg_name_matches_prefix _s1080_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1596_ _) =>
+ (match (string_drop _s1595_ _s1596_) with
+ | _s1597_ =>
+ (reg_name_matches_prefix _s1597_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1081_ _)) =>
- let p0_ := string_drop _s1080_ _s1081_ in
- if ((generic_eq p0_ "")) then Some ((rs1, rs2))
- else None
- | _ => None
- end)
- : option ((mword 5 * mword 5)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s1598_ _) =>
+ let p0_ := string_drop _s1597_ _s1598_ in
+ if generic_eq p0_ "" then Some (rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- else returnm (None : option ((mword 5 * mword 5))))
+ else returnm None)
: M (option ((mword 5 * mword 5))).
-Definition _s1059_ (_s1060_ : string)
-: M (option ((mword 4 * mword 4))) :=
-
- let _s1061_ := _s1060_ in
- (if ((string_startswith _s1061_ "fence.tso")) then
- (match (string_drop _s1061_ (projT1 (string_length "fence.tso"))) with
- | _s1062_ =>
- (spc_matches_prefix _s1062_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1576_ (_s1577_ : string) : M (option ((mword 4 * mword 4))) :=
+ let _s1578_ := _s1577_ in
+ (if string_startswith _s1578_ "fence.tso" then
+ (match (string_drop _s1578_ (projT1 (string_length "fence.tso"))) with
+ | _s1579_ =>
+ (spc_matches_prefix _s1579_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1063_ _)) =>
- (match (string_drop _s1062_ _s1063_) with
- | _s1064_ =>
- (fence_bits_matches_prefix _s1064_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1580_ _) =>
+ (match (string_drop _s1579_ _s1580_) with
+ | _s1581_ =>
+ (fence_bits_matches_prefix _s1581_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((pred, existT _ _s1065_ _)) =>
- (match (string_drop _s1064_ _s1065_) with
- | _s1066_ =>
- (sep_matches_prefix _s1066_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (pred, existT _ _s1582_ _) =>
+ (match (string_drop _s1581_ _s1582_) with
+ | _s1583_ =>
+ (sep_matches_prefix _s1583_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1067_ _)) =>
- (match (string_drop _s1066_ _s1067_) with
- | _s1068_ =>
- (fence_bits_matches_prefix _s1068_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1584_ _) =>
+ (match (string_drop _s1583_ _s1584_) with
+ | _s1585_ =>
+ (fence_bits_matches_prefix _s1585_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((succ, existT _ _s1069_ _)) =>
- let p0_ := string_drop _s1068_ _s1069_ in
- if ((generic_eq p0_ "")) then Some ((pred, succ))
- else None
- | _ => None
- end)
- : option ((mword 4 * mword 4)))
+ returnm (match w__3 with
+ | Some (succ, existT _ _s1586_ _) =>
+ let p0_ := string_drop _s1585_ _s1586_ in
+ if generic_eq p0_ "" then Some (pred, succ)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- else returnm (None : option ((mword 4 * mword 4))))
+ else returnm None)
: M (option ((mword 4 * mword 4))).
-Definition _s1047_ (_s1048_ : string)
-: M (option ((mword 4 * mword 4))) :=
-
- let _s1049_ := _s1048_ in
- (if ((string_startswith _s1049_ "fence")) then
- (match (string_drop _s1049_ (projT1 (string_length "fence"))) with
- | _s1050_ =>
- (spc_matches_prefix _s1050_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1564_ (_s1565_ : string) : M (option ((mword 4 * mword 4))) :=
+ let _s1566_ := _s1565_ in
+ (if string_startswith _s1566_ "fence" then
+ (match (string_drop _s1566_ (projT1 (string_length "fence"))) with
+ | _s1567_ =>
+ (spc_matches_prefix _s1567_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1051_ _)) =>
- (match (string_drop _s1050_ _s1051_) with
- | _s1052_ =>
- (fence_bits_matches_prefix _s1052_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1568_ _) =>
+ (match (string_drop _s1567_ _s1568_) with
+ | _s1569_ =>
+ (fence_bits_matches_prefix _s1569_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((pred, existT _ _s1053_ _)) =>
- (match (string_drop _s1052_ _s1053_) with
- | _s1054_ =>
- (sep_matches_prefix _s1054_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (pred, existT _ _s1570_ _) =>
+ (match (string_drop _s1569_ _s1570_) with
+ | _s1571_ =>
+ (sep_matches_prefix _s1571_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1055_ _)) =>
- (match (string_drop _s1054_ _s1055_) with
- | _s1056_ =>
- (fence_bits_matches_prefix _s1056_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1572_ _) =>
+ (match (string_drop _s1571_ _s1572_) with
+ | _s1573_ =>
+ (fence_bits_matches_prefix _s1573_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((succ, existT _ _s1057_ _)) =>
- let p0_ := string_drop _s1056_ _s1057_ in
- if ((generic_eq p0_ "")) then Some ((pred, succ))
- else None
- | _ => None
- end)
- : option ((mword 4 * mword 4)))
+ returnm (match w__3 with
+ | Some (succ, existT _ _s1574_ _) =>
+ let p0_ := string_drop _s1573_ _s1574_ in
+ if generic_eq p0_ "" then Some (pred, succ)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- else returnm (None : option ((mword 4 * mword 4))))
+ else returnm None)
: M (option ((mword 4 * mword 4))).
-Definition _s1030_ (_s1031_ : string)
-: M (option ((sopw * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1031_ with
- | _s1032_ =>
- (shiftiwop_mnemonic_matches_prefix _s1032_) >>= fun w__0 : option ((sopw * {n : Z & ArithFact (n >=
+Definition _s1547_ (_s1548_ : string) : M (option ((sopw * mword 5 * mword 5 * mword 5))) :=
+ (match _s1548_ with
+ | _s1549_ =>
+ (shiftiwop_mnemonic_matches_prefix _s1549_) >>= fun w__0 : option ((sopw * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1033_ _)) =>
- (match (string_drop _s1032_ _s1033_) with
- | _s1034_ =>
- (spc_matches_prefix _s1034_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s1550_ _) =>
+ (match (string_drop _s1549_ _s1550_) with
+ | _s1551_ =>
+ (spc_matches_prefix _s1551_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1035_ _)) =>
- (match (string_drop _s1034_ _s1035_) with
- | _s1036_ =>
- (reg_name_matches_prefix _s1036_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1552_ _) =>
+ (match (string_drop _s1551_ _s1552_) with
+ | _s1553_ =>
+ (reg_name_matches_prefix _s1553_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1037_ _)) =>
- (match (string_drop _s1036_ _s1037_) with
- | _s1038_ =>
- (sep_matches_prefix _s1038_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1554_ _) =>
+ (match (string_drop _s1553_ _s1554_) with
+ | _s1555_ =>
+ (sep_matches_prefix _s1555_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1039_ _)) =>
- (match (string_drop _s1038_ _s1039_) with
- | _s1040_ =>
- (reg_name_matches_prefix _s1040_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1556_ _) =>
+ (match (string_drop _s1555_ _s1556_) with
+ | _s1557_ =>
+ (reg_name_matches_prefix _s1557_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1041_ _)) =>
- (match (string_drop _s1040_ _s1041_) with
- | _s1042_ =>
- (sep_matches_prefix _s1042_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1558_ _) =>
+ (match (string_drop _s1557_ _s1558_) with
+ | _s1559_ =>
+ (sep_matches_prefix _s1559_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s1043_ _)) =>
- match (string_drop _s1042_ _s1043_) with
- | _s1044_ =>
- match (hex_bits_5_matches_prefix
- _s1044_) with
- | Some ((shamt, existT _ _s1045_ _)) =>
- let p0_ :=
- string_drop _s1044_ _s1045_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sopw * mword 5 * mword 5 * mword 5)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s1560_ _) =>
+ match (string_drop _s1559_ _s1560_) with
+ | _s1561_ =>
+ match (hex_bits_5_matches_prefix
+ _s1561_) with
+ | Some (shamt, existT _ _s1562_ _) =>
+ let p0_ :=
+ string_drop _s1561_ _s1562_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5))).
-Definition _s1013_ (_s1014_ : string)
-: M (option ((ropw * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1014_ with
- | _s1015_ =>
- (rtypew_mnemonic_matches_prefix _s1015_) >>= fun w__0 : option ((ropw * {n : Z & ArithFact (n >=
+Definition _s1530_ (_s1531_ : string) : M (option ((ropw * mword 5 * mword 5 * mword 5))) :=
+ (match _s1531_ with
+ | _s1532_ =>
+ (rtypew_mnemonic_matches_prefix _s1532_) >>= fun w__0 : option ((ropw * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1016_ _)) =>
- (match (string_drop _s1015_ _s1016_) with
- | _s1017_ =>
- (spc_matches_prefix _s1017_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s1533_ _) =>
+ (match (string_drop _s1532_ _s1533_) with
+ | _s1534_ =>
+ (spc_matches_prefix _s1534_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1018_ _)) =>
- (match (string_drop _s1017_ _s1018_) with
- | _s1019_ =>
- (reg_name_matches_prefix _s1019_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1535_ _) =>
+ (match (string_drop _s1534_ _s1535_) with
+ | _s1536_ =>
+ (reg_name_matches_prefix _s1536_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1020_ _)) =>
- (match (string_drop _s1019_ _s1020_) with
- | _s1021_ =>
- (sep_matches_prefix _s1021_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1537_ _) =>
+ (match (string_drop _s1536_ _s1537_) with
+ | _s1538_ =>
+ (sep_matches_prefix _s1538_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1022_ _)) =>
- (match (string_drop _s1021_ _s1022_) with
- | _s1023_ =>
- (reg_name_matches_prefix _s1023_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1539_ _) =>
+ (match (string_drop _s1538_ _s1539_) with
+ | _s1540_ =>
+ (reg_name_matches_prefix _s1540_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1024_ _)) =>
- (match (string_drop _s1023_ _s1024_) with
- | _s1025_ =>
- (sep_matches_prefix _s1025_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1541_ _) =>
+ (match (string_drop _s1540_ _s1541_) with
+ | _s1542_ =>
+ (sep_matches_prefix _s1542_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1026_ _)) =>
- (match (string_drop _s1025_ _s1026_) with
- | _s1027_ =>
- (reg_name_matches_prefix _s1027_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1543_ _) =>
+ (match (string_drop _s1542_ _s1543_) with
+ | _s1544_ =>
+ (reg_name_matches_prefix _s1544_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s1028_ _)) =>
- let p0_ :=
- string_drop _s1027_ _s1028_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((ropw * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s1545_ _) =>
+ let p0_ :=
+ string_drop _s1544_ _s1545_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5))).
-Definition _s996_ (_s997_ : string)
-: M (option ((sop * mword 5 * mword 5 * mword 5))) :=
-
- (match _s997_ with
- | _s998_ =>
- (shiftw_mnemonic_matches_prefix _s998_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=
+Definition _s1513_ (_s1514_ : string) : M (option ((sop * mword 5 * mword 5 * mword 5))) :=
+ (match _s1514_ with
+ | _s1515_ =>
+ (shiftw_mnemonic_matches_prefix _s1515_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s999_ _)) =>
- (match (string_drop _s998_ _s999_) with
- | _s1000_ =>
- (spc_matches_prefix _s1000_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s1516_ _) =>
+ (match (string_drop _s1515_ _s1516_) with
+ | _s1517_ =>
+ (spc_matches_prefix _s1517_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1001_ _)) =>
- (match (string_drop _s1000_ _s1001_) with
- | _s1002_ =>
- (reg_name_matches_prefix _s1002_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1518_ _) =>
+ (match (string_drop _s1517_ _s1518_) with
+ | _s1519_ =>
+ (reg_name_matches_prefix _s1519_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1003_ _)) =>
- (match (string_drop _s1002_ _s1003_) with
- | _s1004_ =>
- (sep_matches_prefix _s1004_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1520_ _) =>
+ (match (string_drop _s1519_ _s1520_) with
+ | _s1521_ =>
+ (sep_matches_prefix _s1521_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1005_ _)) =>
- (match (string_drop _s1004_ _s1005_) with
- | _s1006_ =>
- (reg_name_matches_prefix _s1006_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1522_ _) =>
+ (match (string_drop _s1521_ _s1522_) with
+ | _s1523_ =>
+ (reg_name_matches_prefix _s1523_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1007_ _)) =>
- (match (string_drop _s1006_ _s1007_) with
- | _s1008_ =>
- (sep_matches_prefix _s1008_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1524_ _) =>
+ (match (string_drop _s1523_ _s1524_) with
+ | _s1525_ =>
+ (sep_matches_prefix _s1525_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s1009_ _)) =>
- match (string_drop _s1008_ _s1009_) with
- | _s1010_ =>
- match (hex_bits_5_matches_prefix
- _s1010_) with
- | Some ((shamt, existT _ _s1011_ _)) =>
- let p0_ :=
- string_drop _s1010_ _s1011_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sop * mword 5 * mword 5 * mword 5)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s1526_ _) =>
+ match (string_drop _s1525_ _s1526_) with
+ | _s1527_ =>
+ match (hex_bits_5_matches_prefix
+ _s1527_) with
+ | Some (shamt, existT _ _s1528_ _) =>
+ let p0_ :=
+ string_drop _s1527_ _s1528_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5))).
-Definition _s980_ (_s981_ : string)
-: M (option ((mword 5 * mword 5 * mword 12))) :=
-
- let _s982_ := _s981_ in
- (if ((string_startswith _s982_ "addiw")) then
- (match (string_drop _s982_ (projT1 (string_length "addiw"))) with
- | _s983_ =>
- (spc_matches_prefix _s983_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1497_ (_s1498_ : string) : M (option ((mword 5 * mword 5 * mword 12))) :=
+ let _s1499_ := _s1498_ in
+ (if string_startswith _s1499_ "addiw" then
+ (match (string_drop _s1499_ (projT1 (string_length "addiw"))) with
+ | _s1500_ =>
+ (spc_matches_prefix _s1500_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s984_ _)) =>
- (match (string_drop _s983_ _s984_) with
- | _s985_ =>
- (reg_name_matches_prefix _s985_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1501_ _) =>
+ (match (string_drop _s1500_ _s1501_) with
+ | _s1502_ =>
+ (reg_name_matches_prefix _s1502_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s986_ _)) =>
- (match (string_drop _s985_ _s986_) with
- | _s987_ =>
- (sep_matches_prefix _s987_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1503_ _) =>
+ (match (string_drop _s1502_ _s1503_) with
+ | _s1504_ =>
+ (sep_matches_prefix _s1504_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s988_ _)) =>
- (match (string_drop _s987_ _s988_) with
- | _s989_ =>
- (reg_name_matches_prefix _s989_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1505_ _) =>
+ (match (string_drop _s1504_ _s1505_) with
+ | _s1506_ =>
+ (reg_name_matches_prefix _s1506_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s990_ _)) =>
- (match (string_drop _s989_ _s990_) with
- | _s991_ =>
- (sep_matches_prefix _s991_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1507_ _) =>
+ (match (string_drop _s1506_ _s1507_) with
+ | _s1508_ =>
+ (sep_matches_prefix _s1508_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s992_ _)) =>
- match (string_drop _s991_ _s992_) with
- | _s993_ =>
- match (hex_bits_12_matches_prefix _s993_) with
- | Some ((imm, existT _ _s994_ _)) =>
- let p0_ := string_drop _s993_ _s994_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rd, rs1, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 12)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s1509_ _) =>
+ match (string_drop _s1508_ _s1509_) with
+ | _s1510_ =>
+ match (hex_bits_12_matches_prefix _s1510_) with
+ | Some (imm, existT _ _s1511_ _) =>
+ let p0_ := string_drop _s1510_ _s1511_ in
+ if generic_eq p0_ "" then
+ Some (rd, rs1, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- else returnm (None : option ((mword 5 * mword 5 * mword 12))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 12))).
-Definition _s952_ (_s953_ : string)
+Definition _s1469_ (_s1470_ : string)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5))) :=
-
- let _s954_ := _s953_ in
- (if ((string_startswith _s954_ "s")) then
- (match (string_drop _s954_ (projT1 (string_length "s"))) with
- | _s955_ =>
- (size_mnemonic_matches_prefix _s955_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s1471_ := _s1470_ in
+ (if string_startswith _s1471_ "s" then
+ (match (string_drop _s1471_ (projT1 (string_length "s"))) with
+ | _s1472_ =>
+ (size_mnemonic_matches_prefix _s1472_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s956_ _)) =>
- (match (string_drop _s955_ _s956_) with
- | _s957_ =>
- (maybe_aq_matches_prefix _s957_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s1473_ _) =>
+ (match (string_drop _s1472_ _s1473_) with
+ | _s1474_ =>
+ (maybe_aq_matches_prefix _s1474_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s958_ _)) =>
- (match (string_drop _s957_ _s958_) with
- | _s959_ =>
- (maybe_rl_matches_prefix _s959_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s1475_ _) =>
+ (match (string_drop _s1474_ _s1475_) with
+ | _s1476_ =>
+ (maybe_rl_matches_prefix _s1476_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s960_ _)) =>
- (match (string_drop _s959_ _s960_) with
- | _s961_ =>
- (spc_matches_prefix _s961_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s1477_ _) =>
+ (match (string_drop _s1476_ _s1477_) with
+ | _s1478_ =>
+ (spc_matches_prefix _s1478_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s962_ _)) =>
- (match (string_drop _s961_ _s962_) with
- | _s963_ =>
- (reg_name_matches_prefix _s963_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1479_ _) =>
+ (match (string_drop _s1478_ _s1479_) with
+ | _s1480_ =>
+ (reg_name_matches_prefix _s1480_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs2, existT _ _s964_ _)) =>
- (match (string_drop _s963_ _s964_) with
- | _s965_ =>
- (sep_matches_prefix _s965_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s1481_ _) =>
+ (match (string_drop _s1480_ _s1481_) with
+ | _s1482_ =>
+ (sep_matches_prefix _s1482_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s966_ _)) =>
- (match (string_drop _s965_ _s966_) with
- | _s967_ =>
- (match (hex_bits_12_matches_prefix _s967_) with
- | Some ((imm, existT _ _s968_ _)) =>
- (match (string_drop _s967_ _s968_) with
- | _s969_ =>
- (opt_spc_matches_prefix _s969_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1483_ _) =>
+ (match (string_drop _s1482_ _s1483_) with
+ | _s1484_ =>
+ (match (hex_bits_12_matches_prefix _s1484_) with
+ | Some (imm, existT _ _s1485_ _) =>
+ (match (string_drop _s1484_ _s1485_) with
+ | _s1486_ =>
+ (opt_spc_matches_prefix _s1486_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s970_ _)) =>
- let _s971_ :=
- string_drop _s969_ _s970_ in
- (if ((string_startswith _s971_
- "(")) then
- (match (string_drop _s971_
+ | Some (tt, existT _ _s1487_ _) =>
+ let _s1488_ :=
+ string_drop _s1486_ _s1487_ in
+ (if string_startswith _s1488_
+ "(" then
+ (match (string_drop _s1488_
(projT1
(string_length
"("))) with
- | _s972_ =>
+ | _s1489_ =>
(opt_spc_matches_prefix
- _s972_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ _s1489_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
| Some
- ((tt, existT _ _s973_ _)) =>
+ (tt, existT _ _s1490_ _) =>
(match (string_drop
- _s972_
- _s973_) with
- | _s974_ =>
+ _s1489_
+ _s1490_) with
+ | _s1491_ =>
(reg_name_matches_prefix
- _s974_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s1491_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((rs1, existT _ _s975_ _)) =>
+ (rs1, existT _ _s1492_ _) =>
(match (string_drop
- _s974_
- _s975_) with
- | _s976_ =>
+ _s1491_
+ _s1492_) with
+ | _s1493_ =>
(opt_spc_matches_prefix
- _s976_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=
+ _s1493_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__9 with
- | Some
- ((tt, existT _ _s977_ _)) =>
- let _s978_ :=
- string_drop
- _s976_
- _s977_ in
- if
- ((string_startswith
- _s978_
- ")"))
- then
- let p0_ :=
- string_drop
- _s978_
- (projT1
- (string_length
- ")")) in
- if
- ((generic_eq
- p0_
- ""))
- then
- Some
- ((size, aq, rl, rs2, imm, rs1))
- else
- None
- else
- None
- | _ =>
- None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm (match w__9 with
+ | Some
+ (tt, existT _ _s1494_ _) =>
+ let _s1495_ :=
+ string_drop
+ _s1493_
+ _s1494_ in
+ if
+ string_startswith
+ _s1495_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s1495_
+ (projT1
+ (string_length
+ ")")) in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (size, aq, rl, rs2, imm, rs1)
+ else
+ None
+ else
+ None
+ | _ =>
+ None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- else
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- else returnm (None : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5))).
-Definition _s922_ (_s923_ : string)
+Definition _s1439_ (_s1440_ : string)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5))) :=
-
- let _s924_ := _s923_ in
- (if ((string_startswith _s924_ "l")) then
- (match (string_drop _s924_ (projT1 (string_length "l"))) with
- | _s925_ =>
- (size_mnemonic_matches_prefix _s925_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s1441_ := _s1440_ in
+ (if string_startswith _s1441_ "l" then
+ (match (string_drop _s1441_ (projT1 (string_length "l"))) with
+ | _s1442_ =>
+ (size_mnemonic_matches_prefix _s1442_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s926_ _)) =>
- (match (string_drop _s925_ _s926_) with
- | _s927_ =>
- (maybe_u_matches_prefix _s927_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s1443_ _) =>
+ (match (string_drop _s1442_ _s1443_) with
+ | _s1444_ =>
+ (maybe_u_matches_prefix _s1444_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((is_unsigned, existT _ _s928_ _)) =>
- (match (string_drop _s927_ _s928_) with
- | _s929_ =>
- (maybe_aq_matches_prefix _s929_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (is_unsigned, existT _ _s1445_ _) =>
+ (match (string_drop _s1444_ _s1445_) with
+ | _s1446_ =>
+ (maybe_aq_matches_prefix _s1446_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((aq, existT _ _s930_ _)) =>
- (match (string_drop _s929_ _s930_) with
- | _s931_ =>
- (maybe_rl_matches_prefix _s931_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s1447_ _) =>
+ (match (string_drop _s1446_ _s1447_) with
+ | _s1448_ =>
+ (maybe_rl_matches_prefix _s1448_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rl, existT _ _s932_ _)) =>
- (match (string_drop _s931_ _s932_) with
- | _s933_ =>
- (spc_matches_prefix _s933_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s1449_ _) =>
+ (match (string_drop _s1448_ _s1449_) with
+ | _s1450_ =>
+ (spc_matches_prefix _s1450_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s934_ _)) =>
- (match (string_drop _s933_ _s934_) with
- | _s935_ =>
- (reg_name_matches_prefix _s935_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1451_ _) =>
+ (match (string_drop _s1450_ _s1451_) with
+ | _s1452_ =>
+ (reg_name_matches_prefix _s1452_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((rd, existT _ _s936_ _)) =>
- (match (string_drop _s935_ _s936_) with
- | _s937_ =>
- (sep_matches_prefix _s937_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1453_ _) =>
+ (match (string_drop _s1452_ _s1453_) with
+ | _s1454_ =>
+ (sep_matches_prefix _s1454_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s938_ _)) =>
- (match (string_drop _s937_ _s938_) with
- | _s939_ =>
+ | Some (tt, existT _ _s1455_ _) =>
+ (match (string_drop _s1454_ _s1455_) with
+ | _s1456_ =>
(match (hex_bits_12_matches_prefix
- _s939_) with
- | Some ((imm, existT _ _s940_ _)) =>
- (match (string_drop _s939_
- _s940_) with
- | _s941_ =>
+ _s1456_) with
+ | Some (imm, existT _ _s1457_ _) =>
+ (match (string_drop _s1456_
+ _s1457_) with
+ | _s1458_ =>
(opt_spc_matches_prefix
- _s941_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ _s1458_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
| Some
- ((tt, existT _ _s942_ _)) =>
- let _s943_ :=
- string_drop _s941_
- _s942_ in
- (if ((string_startswith
- _s943_ "("))
- then
+ (tt, existT _ _s1459_ _) =>
+ let _s1460_ :=
+ string_drop _s1458_
+ _s1459_ in
+ (if string_startswith
+ _s1460_ "(" then
(match (string_drop
- _s943_
+ _s1460_
(projT1
(string_length
"("))) with
- | _s944_ =>
+ | _s1461_ =>
(opt_spc_matches_prefix
- _s944_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=
+ _s1461_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((tt, existT _ _s945_ _)) =>
+ (tt, existT _ _s1462_ _) =>
(match (string_drop
- _s944_
- _s945_) with
- | _s946_ =>
+ _s1461_
+ _s1462_) with
+ | _s1463_ =>
(reg_name_matches_prefix
- _s946_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s1463_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__9 with
| Some
- ((rs1, existT _ _s947_ _)) =>
+ (rs1, existT _ _s1464_ _) =>
(match (string_drop
- _s946_
- _s947_) with
- | _s948_ =>
+ _s1463_
+ _s1464_) with
+ | _s1465_ =>
(opt_spc_matches_prefix
- _s948_) >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=
+ _s1465_) >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__10 with
- | Some
- ((tt, existT _ _s949_ _)) =>
- let _s950_ :=
- string_drop
- _s948_
- _s949_ in
- if
- ((string_startswith
- _s950_
- ")"))
- then
- let p0_ :=
- string_drop
- _s950_
- (projT1
- (string_length
- ")")) in
- if
- ((generic_eq
- p0_
- ""))
- then
- Some
- ((size, is_unsigned, aq, rl, rd, imm, rs1))
- else
- None
- else
- None
- | _ =>
- None
- end)
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm (match w__10 with
+ | Some
+ (tt, existT _ _s1466_ _) =>
+ let _s1467_ :=
+ string_drop
+ _s1465_
+ _s1466_ in
+ if
+ string_startswith
+ _s1467_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s1467_
+ (projT1
+ (string_length
+ ")")) in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (size, is_unsigned, aq, rl, rd, imm, rs1)
+ else
+ None
+ else
+ None
+ | _ =>
+ None
+ end)
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- else
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- else returnm (None : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5))).
-Definition _s905_ (_s906_ : string)
-: M (option ((rop * mword 5 * mword 5 * mword 5))) :=
-
- (match _s906_ with
- | _s907_ =>
- (rtype_mnemonic_matches_prefix _s907_) >>= fun w__0 : option ((rop * {n : Z & ArithFact (n >=
+Definition _s1422_ (_s1423_ : string) : M (option ((rop * mword 5 * mword 5 * mword 5))) :=
+ (match _s1423_ with
+ | _s1424_ =>
+ (rtype_mnemonic_matches_prefix _s1424_) >>= fun w__0 : option ((rop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s908_ _)) =>
- (match (string_drop _s907_ _s908_) with
- | _s909_ =>
- (spc_matches_prefix _s909_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ | Some (op, existT _ _s1425_ _) =>
+ (match (string_drop _s1424_ _s1425_) with
+ | _s1426_ =>
+ (spc_matches_prefix _s1426_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s910_ _)) =>
- (match (string_drop _s909_ _s910_) with
- | _s911_ =>
- (reg_name_matches_prefix _s911_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1427_ _) =>
+ (match (string_drop _s1426_ _s1427_) with
+ | _s1428_ =>
+ (reg_name_matches_prefix _s1428_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s912_ _)) =>
- (match (string_drop _s911_ _s912_) with
- | _s913_ =>
- (sep_matches_prefix _s913_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1429_ _) =>
+ (match (string_drop _s1428_ _s1429_) with
+ | _s1430_ =>
+ (sep_matches_prefix _s1430_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s914_ _)) =>
- (match (string_drop _s913_ _s914_) with
- | _s915_ =>
- (reg_name_matches_prefix _s915_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1431_ _) =>
+ (match (string_drop _s1430_ _s1431_) with
+ | _s1432_ =>
+ (reg_name_matches_prefix _s1432_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s916_ _)) =>
- (match (string_drop _s915_ _s916_) with
- | _s917_ =>
- (sep_matches_prefix _s917_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1433_ _) =>
+ (match (string_drop _s1432_ _s1433_) with
+ | _s1434_ =>
+ (sep_matches_prefix _s1434_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s918_ _)) =>
- (match (string_drop _s917_ _s918_) with
- | _s919_ =>
- (reg_name_matches_prefix _s919_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1435_ _) =>
+ (match (string_drop _s1434_ _s1435_) with
+ | _s1436_ =>
+ (reg_name_matches_prefix _s1436_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s920_ _)) =>
- let p0_ :=
- string_drop _s919_ _s920_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((rop * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s1437_ _) =>
+ let p0_ :=
+ string_drop _s1436_ _s1437_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5))).
-Definition _s888_ (_s889_ : string)
-: M (option ((sop * mword 5 * mword 5 * mword 6))) :=
-
- (match _s889_ with
- | _s890_ =>
- (shiftiop_mnemonic_matches_prefix _s890_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=
+Definition _s1405_ (_s1406_ : string) : M (option ((sop * mword 5 * mword 5 * mword 6))) :=
+ (match _s1406_ with
+ | _s1407_ =>
+ (shiftiop_mnemonic_matches_prefix _s1407_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s891_ _)) =>
- (match (string_drop _s890_ _s891_) with
- | _s892_ =>
- (spc_matches_prefix _s892_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ | Some (op, existT _ _s1408_ _) =>
+ (match (string_drop _s1407_ _s1408_) with
+ | _s1409_ =>
+ (spc_matches_prefix _s1409_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s893_ _)) =>
- (match (string_drop _s892_ _s893_) with
- | _s894_ =>
- (reg_name_matches_prefix _s894_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1410_ _) =>
+ (match (string_drop _s1409_ _s1410_) with
+ | _s1411_ =>
+ (reg_name_matches_prefix _s1411_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s895_ _)) =>
- (match (string_drop _s894_ _s895_) with
- | _s896_ =>
- (sep_matches_prefix _s896_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1412_ _) =>
+ (match (string_drop _s1411_ _s1412_) with
+ | _s1413_ =>
+ (sep_matches_prefix _s1413_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s897_ _)) =>
- (match (string_drop _s896_ _s897_) with
- | _s898_ =>
- (reg_name_matches_prefix _s898_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1414_ _) =>
+ (match (string_drop _s1413_ _s1414_) with
+ | _s1415_ =>
+ (reg_name_matches_prefix _s1415_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s899_ _)) =>
- (match (string_drop _s898_ _s899_) with
- | _s900_ =>
- (sep_matches_prefix _s900_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1416_ _) =>
+ (match (string_drop _s1415_ _s1416_) with
+ | _s1417_ =>
+ (sep_matches_prefix _s1417_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s901_ _)) =>
- match (string_drop _s900_ _s901_) with
- | _s902_ =>
- match (hex_bits_6_matches_prefix
- _s902_) with
- | Some ((shamt, existT _ _s903_ _)) =>
- let p0_ :=
- string_drop _s902_ _s903_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sop * mword 5 * mword 5 * mword 6)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s1418_ _) =>
+ match (string_drop _s1417_ _s1418_) with
+ | _s1419_ =>
+ match (hex_bits_6_matches_prefix
+ _s1419_) with
+ | Some (shamt, existT _ _s1420_ _) =>
+ let p0_ :=
+ string_drop _s1419_ _s1420_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ =>
- returnm (None
- : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ =>
- returnm (None : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6))).
-Definition _s871_ (_s872_ : string)
-: M (option ((iop * mword 5 * mword 5 * mword 12))) :=
-
- (match _s872_ with
- | _s873_ =>
- (itype_mnemonic_matches_prefix _s873_) >>= fun w__0 : option ((iop * {n : Z & ArithFact (n >=
+Definition _s1388_ (_s1389_ : string) : M (option ((iop * mword 5 * mword 5 * mword 12))) :=
+ (match _s1389_ with
+ | _s1390_ =>
+ (itype_mnemonic_matches_prefix _s1390_) >>= fun w__0 : option ((iop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s874_ _)) =>
- (match (string_drop _s873_ _s874_) with
- | _s875_ =>
- (spc_matches_prefix _s875_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ | Some (op, existT _ _s1391_ _) =>
+ (match (string_drop _s1390_ _s1391_) with
+ | _s1392_ =>
+ (spc_matches_prefix _s1392_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s876_ _)) =>
- (match (string_drop _s875_ _s876_) with
- | _s877_ =>
- (reg_name_matches_prefix _s877_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1393_ _) =>
+ (match (string_drop _s1392_ _s1393_) with
+ | _s1394_ =>
+ (reg_name_matches_prefix _s1394_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s878_ _)) =>
- (match (string_drop _s877_ _s878_) with
- | _s879_ =>
- (sep_matches_prefix _s879_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1395_ _) =>
+ (match (string_drop _s1394_ _s1395_) with
+ | _s1396_ =>
+ (sep_matches_prefix _s1396_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s880_ _)) =>
- (match (string_drop _s879_ _s880_) with
- | _s881_ =>
- (reg_name_matches_prefix _s881_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1397_ _) =>
+ (match (string_drop _s1396_ _s1397_) with
+ | _s1398_ =>
+ (reg_name_matches_prefix _s1398_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s882_ _)) =>
- (match (string_drop _s881_ _s882_) with
- | _s883_ =>
- (sep_matches_prefix _s883_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1399_ _) =>
+ (match (string_drop _s1398_ _s1399_) with
+ | _s1400_ =>
+ (sep_matches_prefix _s1400_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s884_ _)) =>
- match (string_drop _s883_ _s884_) with
- | _s885_ =>
- match (hex_bits_12_matches_prefix
- _s885_) with
- | Some ((imm, existT _ _s886_ _)) =>
- let p0_ :=
- string_drop _s885_ _s886_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((iop * mword 5 * mword 5 * mword 12)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s1401_ _) =>
+ match (string_drop _s1400_ _s1401_) with
+ | _s1402_ =>
+ match (hex_bits_12_matches_prefix
+ _s1402_) with
+ | Some (imm, existT _ _s1403_ _) =>
+ let p0_ :=
+ string_drop _s1402_ _s1403_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12))).
-Definition _s854_ (_s855_ : string)
-: M (option ((bop * mword 5 * mword 5 * mword 13))) :=
-
- (match _s855_ with
- | _s856_ =>
- (btype_mnemonic_matches_prefix _s856_) >>= fun w__0 : option ((bop * {n : Z & ArithFact (n >=
+Definition _s1371_ (_s1372_ : string) : M (option ((bop * mword 5 * mword 5 * mword 13))) :=
+ (match _s1372_ with
+ | _s1373_ =>
+ (btype_mnemonic_matches_prefix _s1373_) >>= fun w__0 : option ((bop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s857_ _)) =>
- (match (string_drop _s856_ _s857_) with
- | _s858_ =>
- (spc_matches_prefix _s858_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ | Some (op, existT _ _s1374_ _) =>
+ (match (string_drop _s1373_ _s1374_) with
+ | _s1375_ =>
+ (spc_matches_prefix _s1375_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s859_ _)) =>
- (match (string_drop _s858_ _s859_) with
- | _s860_ =>
- (reg_name_matches_prefix _s860_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1376_ _) =>
+ (match (string_drop _s1375_ _s1376_) with
+ | _s1377_ =>
+ (reg_name_matches_prefix _s1377_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rs1, existT _ _s861_ _)) =>
- (match (string_drop _s860_ _s861_) with
- | _s862_ =>
- (sep_matches_prefix _s862_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1378_ _) =>
+ (match (string_drop _s1377_ _s1378_) with
+ | _s1379_ =>
+ (sep_matches_prefix _s1379_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s863_ _)) =>
- (match (string_drop _s862_ _s863_) with
- | _s864_ =>
- (reg_name_matches_prefix _s864_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1380_ _) =>
+ (match (string_drop _s1379_ _s1380_) with
+ | _s1381_ =>
+ (reg_name_matches_prefix _s1381_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs2, existT _ _s865_ _)) =>
- (match (string_drop _s864_ _s865_) with
- | _s866_ =>
- (sep_matches_prefix _s866_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s1382_ _) =>
+ (match (string_drop _s1381_ _s1382_) with
+ | _s1383_ =>
+ (sep_matches_prefix _s1383_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s867_ _)) =>
- match (string_drop _s866_ _s867_) with
- | _s868_ =>
- match (hex_bits_13_matches_prefix
- _s868_) with
- | Some ((imm, existT _ _s869_ _)) =>
- let p0_ :=
- string_drop _s868_ _s869_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rs1, rs2, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((bop * mword 5 * mword 5 * mword 13)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s1384_ _) =>
+ match (string_drop _s1383_ _s1384_) with
+ | _s1385_ =>
+ match (hex_bits_13_matches_prefix
+ _s1385_) with
+ | Some (imm, existT _ _s1386_ _) =>
+ let p0_ :=
+ string_drop _s1385_ _s1386_ in
+ if generic_eq p0_ "" then
+ Some (op, rs1, rs2, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ =>
- returnm (None
- : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ =>
- returnm (None : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ => returnm (None : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ => returnm (None : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ => returnm (None : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13))).
-Definition _s838_ (_s839_ : string)
-: M (option ((mword 5 * mword 5 * mword 12))) :=
-
- let _s840_ := _s839_ in
- (if ((string_startswith _s840_ "jalr")) then
- (match (string_drop _s840_ (projT1 (string_length "jalr"))) with
- | _s841_ =>
- (spc_matches_prefix _s841_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1355_ (_s1356_ : string) : M (option ((mword 5 * mword 5 * mword 12))) :=
+ let _s1357_ := _s1356_ in
+ (if string_startswith _s1357_ "jalr" then
+ (match (string_drop _s1357_ (projT1 (string_length "jalr"))) with
+ | _s1358_ =>
+ (spc_matches_prefix _s1358_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s842_ _)) =>
- (match (string_drop _s841_ _s842_) with
- | _s843_ =>
- (reg_name_matches_prefix _s843_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1359_ _) =>
+ (match (string_drop _s1358_ _s1359_) with
+ | _s1360_ =>
+ (reg_name_matches_prefix _s1360_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s844_ _)) =>
- (match (string_drop _s843_ _s844_) with
- | _s845_ =>
- (sep_matches_prefix _s845_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1361_ _) =>
+ (match (string_drop _s1360_ _s1361_) with
+ | _s1362_ =>
+ (sep_matches_prefix _s1362_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s846_ _)) =>
- (match (string_drop _s845_ _s846_) with
- | _s847_ =>
- (reg_name_matches_prefix _s847_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1363_ _) =>
+ (match (string_drop _s1362_ _s1363_) with
+ | _s1364_ =>
+ (reg_name_matches_prefix _s1364_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s848_ _)) =>
- (match (string_drop _s847_ _s848_) with
- | _s849_ =>
- (sep_matches_prefix _s849_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s1365_ _) =>
+ (match (string_drop _s1364_ _s1365_) with
+ | _s1366_ =>
+ (sep_matches_prefix _s1366_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s850_ _)) =>
- match (string_drop _s849_ _s850_) with
- | _s851_ =>
- match (hex_bits_12_matches_prefix _s851_) with
- | Some ((imm, existT _ _s852_ _)) =>
- let p0_ := string_drop _s851_ _s852_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rd, rs1, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 12)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s1367_ _) =>
+ match (string_drop _s1366_ _s1367_) with
+ | _s1368_ =>
+ match (hex_bits_12_matches_prefix _s1368_) with
+ | Some (imm, existT _ _s1369_ _) =>
+ let p0_ := string_drop _s1368_ _s1369_ in
+ if generic_eq p0_ "" then
+ Some (rd, rs1, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- else returnm (None : option ((mword 5 * mword 5 * mword 12))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 12))).
-Definition _s826_ (_s827_ : string)
-: M (option ((mword 5 * mword 21))) :=
-
- let _s828_ := _s827_ in
- (if ((string_startswith _s828_ "jal")) then
- (match (string_drop _s828_ (projT1 (string_length "jal"))) with
- | _s829_ =>
- (spc_matches_prefix _s829_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s1343_ (_s1344_ : string) : M (option ((mword 5 * mword 21))) :=
+ let _s1345_ := _s1344_ in
+ (if string_startswith _s1345_ "jal" then
+ (match (string_drop _s1345_ (projT1 (string_length "jal"))) with
+ | _s1346_ =>
+ (spc_matches_prefix _s1346_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s830_ _)) =>
- (match (string_drop _s829_ _s830_) with
- | _s831_ =>
- (reg_name_matches_prefix _s831_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1347_ _) =>
+ (match (string_drop _s1346_ _s1347_) with
+ | _s1348_ =>
+ (reg_name_matches_prefix _s1348_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s832_ _)) =>
- (match (string_drop _s831_ _s832_) with
- | _s833_ =>
- (sep_matches_prefix _s833_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1349_ _) =>
+ (match (string_drop _s1348_ _s1349_) with
+ | _s1350_ =>
+ (sep_matches_prefix _s1350_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s834_ _)) =>
- match (string_drop _s833_ _s834_) with
- | _s835_ =>
- match (hex_bits_21_matches_prefix _s835_) with
- | Some ((imm, existT _ _s836_ _)) =>
- let p0_ := string_drop _s835_ _s836_ in
- if ((generic_eq p0_ "")) then Some ((rd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 21)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s1351_ _) =>
+ match (string_drop _s1350_ _s1351_) with
+ | _s1352_ =>
+ match (hex_bits_21_matches_prefix _s1352_) with
+ | Some (imm, existT _ _s1353_ _) =>
+ let p0_ := string_drop _s1352_ _s1353_ in
+ if generic_eq p0_ "" then Some (rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 21)))
- | _ => returnm (None : option ((mword 5 * mword 21)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 21)))
end)
: M (option ((mword 5 * mword 21)))
- | _ => returnm (None : option ((mword 5 * mword 21)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 21)))
end)
: M (option ((mword 5 * mword 21)))
- else returnm (None : option ((mword 5 * mword 21))))
+ else returnm None)
: M (option ((mword 5 * mword 21))).
-Definition _s813_ (_s814_ : string)
-: M (option ((uop * mword 5 * mword 20))) :=
-
- (match _s814_ with
- | _s815_ =>
- (utype_mnemonic_matches_prefix _s815_) >>= fun w__0 : option ((uop * {n : Z & ArithFact (n >=
+Definition _s1330_ (_s1331_ : string) : M (option ((uop * mword 5 * mword 20))) :=
+ (match _s1331_ with
+ | _s1332_ =>
+ (utype_mnemonic_matches_prefix _s1332_) >>= fun w__0 : option ((uop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s816_ _)) =>
- (match (string_drop _s815_ _s816_) with
- | _s817_ =>
- (spc_matches_prefix _s817_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+ | Some (op, existT _ _s1333_ _) =>
+ (match (string_drop _s1332_ _s1333_) with
+ | _s1334_ =>
+ (spc_matches_prefix _s1334_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s818_ _)) =>
- (match (string_drop _s817_ _s818_) with
- | _s819_ =>
- (reg_name_matches_prefix _s819_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s1335_ _) =>
+ (match (string_drop _s1334_ _s1335_) with
+ | _s1336_ =>
+ (reg_name_matches_prefix _s1336_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s820_ _)) =>
- (match (string_drop _s819_ _s820_) with
- | _s821_ =>
- (sep_matches_prefix _s821_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s1337_ _) =>
+ (match (string_drop _s1336_ _s1337_) with
+ | _s1338_ =>
+ (sep_matches_prefix _s1338_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((tt, existT _ _s822_ _)) =>
- match (string_drop _s821_ _s822_) with
- | _s823_ =>
- match (hex_bits_20_matches_prefix _s823_) with
- | Some ((imm, existT _ _s824_ _)) =>
- let p0_ := string_drop _s823_ _s824_ in
- if ((generic_eq p0_ "")) then Some ((op, rd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((uop * mword 5 * mword 20)))
+ returnm (match w__3 with
+ | Some (tt, existT _ _s1339_ _) =>
+ match (string_drop _s1338_ _s1339_) with
+ | _s1340_ =>
+ match (hex_bits_20_matches_prefix _s1340_) with
+ | Some (imm, existT _ _s1341_ _) =>
+ let p0_ := string_drop _s1340_ _s1341_ in
+ if generic_eq p0_ "" then Some (op, rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((uop * mword 5 * mword 20)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20)))
end)
: M (option ((uop * mword 5 * mword 20)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20)))
end)
: M (option ((uop * mword 5 * mword 20)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20)))
end)
: M (option ((uop * mword 5 * mword 20))).
-Definition assembly_backwards (arg_ : string)
-: M (ast) :=
-
- let _s825_ := arg_ in
- (_s813_ _s825_) >>= fun w__0 : option ((uop * mword 5 * mword 20)) =>
- (if ((match w__0 with | Some ((op, rd, imm)) => true | _ => false end)) then
- (_s813_ _s825_) >>= fun w__1 : option ((uop * mword 5 * mword 20)) =>
+Definition assembly_backwards (arg_ : string) : M (ast) :=
+ let _s1342_ := arg_ in
+ (_s1330_ _s1342_) >>= fun w__0 : option ((uop * mword 5 * mword 20)) =>
+ (if match w__0 with | Some (op, rd, imm) => true | _ => false end then
+ (_s1330_ _s1342_) >>= fun w__1 : option ((uop * mword 5 * mword 20)) =>
(match w__1 with
- | Some ((op, rd, imm)) => returnm ((UTYPE ((imm, rd, op))) : ast )
+ | Some (op, rd, imm) => returnm (UTYPE (imm, rd, op))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s826_ _s825_) >>= fun w__4 : option ((mword 5 * mword 21)) =>
- (if ((match w__4 with | Some ((rd, imm)) => true | _ => false end)) then
- (_s826_ _s825_) >>= fun w__5 : option ((mword 5 * mword 21)) =>
+ (_s1343_ _s1342_) >>= fun w__4 : option ((mword 5 * mword 21)) =>
+ (if match w__4 with | Some (rd, imm) => true | _ => false end then
+ (_s1343_ _s1342_) >>= fun w__5 : option ((mword 5 * mword 21)) =>
(match w__5 with
- | Some ((rd, imm)) => returnm ((RISCV_JAL ((imm, rd))) : ast )
+ | Some (rd, imm) => returnm (RISCV_JAL (imm, rd))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s838_ _s825_) >>= fun w__8 : option ((mword 5 * mword 5 * mword 12)) =>
- (if ((match w__8 with | Some ((rd, rs1, imm)) => true | _ => false end)) then
- (_s838_ _s825_) >>= fun w__9 : option ((mword 5 * mword 5 * mword 12)) =>
+ (_s1355_ _s1342_) >>= fun w__8 : option ((mword 5 * mword 5 * mword 12)) =>
+ (if match w__8 with | Some (rd, rs1, imm) => true | _ => false end then
+ (_s1355_ _s1342_) >>= fun w__9 : option ((mword 5 * mword 5 * mword 12)) =>
(match w__9 with
- | Some ((rd, rs1, imm)) => returnm ((RISCV_JALR ((imm, rs1, rd))) : ast )
+ | Some (rd, rs1, imm) => returnm (RISCV_JALR (imm, rs1, rd))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s854_ _s825_) >>= fun w__12 : option ((bop * mword 5 * mword 5 * mword 13)) =>
- (if ((match w__12 with | Some ((op, rs1, rs2, imm)) => true | _ => false end)) then
- (_s854_ _s825_) >>= fun w__13 : option ((bop * mword 5 * mword 5 * mword 13)) =>
+ (_s1371_ _s1342_) >>= fun w__12 : option ((bop * mword 5 * mword 5 * mword 13)) =>
+ (if match w__12 with | Some (op, rs1, rs2, imm) => true | _ => false end then
+ (_s1371_ _s1342_) >>= fun w__13 : option ((bop * mword 5 * mword 5 * mword 13)) =>
(match w__13 with
- | Some ((op, rs1, rs2, imm)) => returnm ((BTYPE ((imm, rs2, rs1, op))) : ast )
+ | Some (op, rs1, rs2, imm) => returnm (BTYPE (imm, rs2, rs1, op))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s871_ _s825_) >>= fun w__16 : option ((iop * mword 5 * mword 5 * mword 12)) =>
- (if ((match w__16 with | Some ((op, rd, rs1, imm)) => true | _ => false end)) then
- (_s871_ _s825_) >>= fun w__17 : option ((iop * mword 5 * mword 5 * mword 12)) =>
+ (_s1388_ _s1342_) >>= fun w__16 : option ((iop * mword 5 * mword 5 * mword 12)) =>
+ (if match w__16 with | Some (op, rd, rs1, imm) => true | _ => false end then
+ (_s1388_ _s1342_) >>= fun w__17 : option ((iop * mword 5 * mword 5 * mword 12)) =>
(match w__17 with
- | Some ((op, rd, rs1, imm)) => returnm ((ITYPE ((imm, rs1, rd, op))) : ast )
+ | Some (op, rd, rs1, imm) => returnm (ITYPE (imm, rs1, rd, op))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s888_ _s825_) >>= fun w__20 : option ((sop * mword 5 * mword 5 * mword 6)) =>
- (if ((match w__20 with | Some ((op, rd, rs1, shamt)) => true | _ => false end))
- then
- (_s888_ _s825_) >>= fun w__21 : option ((sop * mword 5 * mword 5 * mword 6)) =>
+ (_s1405_ _s1342_) >>= fun w__20 : option ((sop * mword 5 * mword 5 * mword 6)) =>
+ (if match w__20 with | Some (op, rd, rs1, shamt) => true | _ => false end then
+ (_s1405_ _s1342_) >>= fun w__21 : option ((sop * mword 5 * mword 5 * mword 6)) =>
(match w__21 with
- | Some ((op, rd, rs1, shamt)) =>
- returnm ((SHIFTIOP ((shamt, rs1, rd, op))) : ast )
+ | Some (op, rd, rs1, shamt) => returnm (SHIFTIOP (shamt, rs1, rd, op))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s905_ _s825_) >>= fun w__24 : option ((rop * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__24 with | Some ((op, rd, rs1, rs2)) => true | _ => false end))
- then
- (_s905_ _s825_) >>= fun w__25 : option ((rop * mword 5 * mword 5 * mword 5)) =>
+ (_s1422_ _s1342_) >>= fun w__24 : option ((rop * mword 5 * mword 5 * mword 5)) =>
+ (if match w__24 with | Some (op, rd, rs1, rs2) => true | _ => false end then
+ (_s1422_ _s1342_) >>= fun w__25 : option ((rop * mword 5 * mword 5 * mword 5)) =>
(match w__25 with
- | Some ((op, rd, rs1, rs2)) =>
- returnm ((RTYPE ((rs2, rs1, rd, op))) : ast )
+ | Some (op, rd, rs1, rs2) => returnm (RTYPE (rs2, rs1, rd, op))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s922_ _s825_) >>= fun w__28 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)) =>
- (if ((match w__28 with
- | Some ((size, is_unsigned, aq, rl, rd, imm, rs1)) => true
- | _ => false
- end)) then
- (_s922_ _s825_) >>= fun w__29 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)) =>
+ (_s1439_ _s1342_) >>= fun w__28 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)) =>
+ (if match w__28 with
+ | Some (size, is_unsigned, aq, rl, rd, imm, rs1) => true
+ | _ => false
+ end then
+ (_s1439_ _s1342_) >>= fun w__29 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)) =>
(match w__29 with
- | Some ((size, is_unsigned, aq, rl, rd, imm, rs1)) =>
- returnm ((LOAD ((imm, rs1, rd, is_unsigned, size, aq, rl))) : ast )
+ | Some (size, is_unsigned, aq, rl, rd, imm, rs1) =>
+ returnm (LOAD (imm, rs1, rd, is_unsigned, size, aq, rl))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s952_ _s825_) >>= fun w__32 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)) =>
- (if ((match w__32 with
- | Some ((size, aq, rl, rs2, imm, rs1)) => true
- | _ => false
- end)) then
- (_s952_ _s825_) >>= fun w__33 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)) =>
+ (_s1469_ _s1342_) >>= fun w__32 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)) =>
+ (if match w__32 with
+ | Some (size, aq, rl, rs2, imm, rs1) => true
+ | _ => false
+ end then
+ (_s1469_ _s1342_) >>= fun w__33 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)) =>
(match w__33 with
- | Some ((size, aq, rl, rs2, imm, rs1)) =>
- returnm ((STORE ((imm, rs2, rs1, size, aq, rl))) : ast )
+ | Some (size, aq, rl, rs2, imm, rs1) =>
+ returnm (STORE (imm, rs2, rs1, size, aq, rl))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s980_ _s825_) >>= fun w__36 : option ((mword 5 * mword 5 * mword 12)) =>
- (if ((match w__36 with
- | Some ((rd, rs1, imm)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s980_ _s825_) >>= fun w__37 : option ((mword 5 * mword 5 * mword 12)) =>
+ (_s1497_ _s1342_) >>= fun w__36 : option ((mword 5 * mword 5 * mword 12)) =>
+ (if match w__36 with
+ | Some (rd, rs1, imm) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s1497_ _s1342_) >>= fun w__37 : option ((mword 5 * mword 5 * mword 12)) =>
(match w__37 with
- | Some ((rd, rs1, imm)) =>
- returnm ((ADDIW ((imm, rs1, rd))) : ast )
+ | Some (rd, rs1, imm) => returnm (ADDIW (imm, rs1, rd))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s996_ _s825_) >>= fun w__40 : option ((sop * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__40 with
- | Some ((op, rd, rs1, shamt)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s996_ _s825_) >>= fun w__41 : option ((sop * mword 5 * mword 5 * mword 5)) =>
+ (_s1513_ _s1342_) >>= fun w__40 : option ((sop * mword 5 * mword 5 * mword 5)) =>
+ (if match w__40 with
+ | Some (op, rd, rs1, shamt) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s1513_ _s1342_) >>= fun w__41 : option ((sop * mword 5 * mword 5 * mword 5)) =>
(match w__41 with
- | Some ((op, rd, rs1, shamt)) =>
- returnm ((SHIFTW ((shamt, rs1, rd, op))) : ast )
+ | Some (op, rd, rs1, shamt) =>
+ returnm (SHIFTW (shamt, rs1, rd, op))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1013_ _s825_) >>= fun w__44 : option ((ropw * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__44 with
- | Some ((op, rd, rs1, rs2)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s1013_ _s825_) >>= fun w__45 : option ((ropw * mword 5 * mword 5 * mword 5)) =>
+ (_s1530_ _s1342_) >>= fun w__44 : option ((ropw * mword 5 * mword 5 * mword 5)) =>
+ (if match w__44 with
+ | Some (op, rd, rs1, rs2) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s1530_ _s1342_) >>= fun w__45 : option ((ropw * mword 5 * mword 5 * mword 5)) =>
(match w__45 with
- | Some ((op, rd, rs1, rs2)) =>
- returnm ((RTYPEW ((rs2, rs1, rd, op))) : ast )
+ | Some (op, rd, rs1, rs2) =>
+ returnm (RTYPEW (rs2, rs1, rd, op))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1030_ _s825_) >>= fun w__48 : option ((sopw * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__48 with
- | Some ((op, rd, rs1, shamt)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s1030_ _s825_) >>= fun w__49 : option ((sopw * mword 5 * mword 5 * mword 5)) =>
+ (_s1547_ _s1342_) >>= fun w__48 : option ((sopw * mword 5 * mword 5 * mword 5)) =>
+ (if match w__48 with
+ | Some (op, rd, rs1, shamt) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s1547_ _s1342_) >>= fun w__49 : option ((sopw * mword 5 * mword 5 * mword 5)) =>
(match w__49 with
- | Some ((op, rd, rs1, shamt)) =>
- returnm ((SHIFTIWOP ((shamt, rs1, rd, op))) : ast )
+ | Some (op, rd, rs1, shamt) =>
+ returnm (SHIFTIWOP (shamt, rs1, rd, op))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1047_ _s825_) >>= fun w__52 : option ((mword 4 * mword 4)) =>
- (if ((match w__52 with
- | Some ((pred, succ)) => true
- | _ => false
- end)) then
- (_s1047_ _s825_) >>= fun w__53 : option ((mword 4 * mword 4)) =>
+ (_s1564_ _s1342_) >>= fun w__52 : option ((mword 4 * mword 4)) =>
+ (if match w__52 with
+ | Some (pred, succ) => true
+ | _ => false
+ end then
+ (_s1564_ _s1342_) >>= fun w__53 : option ((mword 4 * mword 4)) =>
(match w__53 with
- | Some ((pred, succ)) =>
- returnm ((FENCE ((pred, succ))) : ast )
+ | Some (pred, succ) => returnm (FENCE (pred, succ))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1059_ _s825_) >>= fun w__56 : option ((mword 4 * mword 4)) =>
- (if ((match w__56 with
- | Some ((pred, succ)) => true
- | _ => false
- end)) then
- (_s1059_ _s825_) >>= fun w__57 : option ((mword 4 * mword 4)) =>
+ (_s1576_ _s1342_) >>= fun w__56 : option ((mword 4 * mword 4)) =>
+ (if match w__56 with
+ | Some (pred, succ) => true
+ | _ => false
+ end then
+ (_s1576_ _s1342_) >>= fun w__57 : option ((mword 4 * mword 4)) =>
(match w__57 with
- | Some ((pred, succ)) =>
- returnm ((FENCE_TSO ((pred, succ))) : ast )
+ | Some (pred, succ) =>
+ returnm (FENCE_TSO (pred, succ))
| _ => exit tt : M (ast)
end)
: M (ast)
- else if ((generic_eq _s825_ "fence.i")) then
- returnm ((FENCEI
- (tt))
- : ast )
- else if ((generic_eq _s825_ "ecall")) then
- returnm ((ECALL
- (tt))
- : ast )
- else if ((generic_eq _s825_ "mret")) then
- returnm ((MRET
- (tt))
- : ast )
- else if ((generic_eq _s825_ "sret")) then
- returnm ((SRET
- (tt))
- : ast )
- else if ((generic_eq _s825_ "ebreak")) then
- returnm ((EBREAK
- (tt))
- : ast )
- else if ((generic_eq _s825_ "wfi")) then
- returnm ((WFI
- (tt))
- : ast )
+ else if generic_eq _s1342_ "fence.i" then
+ returnm (FENCEI tt)
+ else if generic_eq _s1342_ "ecall" then
+ returnm (ECALL tt)
+ else if generic_eq _s1342_ "mret" then
+ returnm (MRET tt)
+ else if generic_eq _s1342_ "sret" then
+ returnm (SRET tt)
+ else if generic_eq _s1342_ "ebreak" then
+ returnm (EBREAK tt)
+ else if generic_eq _s1342_ "wfi" then returnm (WFI tt)
else
- (_s1071_ _s825_) >>= fun w__60 : option ((mword 5 * mword 5)) =>
- (if ((match w__60 with
- | Some ((rs1, rs2)) => true
- | _ => false
- end)) then
- (_s1071_ _s825_) >>= fun w__61 : option ((mword 5 * mword 5)) =>
+ (_s1588_ _s1342_) >>= fun w__60 : option ((mword 5 * mword 5)) =>
+ (if match w__60 with
+ | Some (rs1, rs2) => true
+ | _ => false
+ end then
+ (_s1588_ _s1342_) >>= fun w__61 : option ((mword 5 * mword 5)) =>
(match w__61 with
- | Some ((rs1, rs2)) =>
- returnm ((SFENCE_VMA ((rs1, rs2))) : ast )
+ | Some (rs1, rs2) =>
+ returnm (SFENCE_VMA (rs1, rs2))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1083_ _s825_) >>= fun w__64 : option ((word_width * bool * bool * mword 5 * mword 5)) =>
- (if ((match w__64 with
- | Some ((size, aq, rl, rd, rs1)) => true
- | _ => false
- end)) then
- (_s1083_ _s825_) >>= fun w__65 : option ((word_width * bool * bool * mword 5 * mword 5)) =>
+ (_s1600_ _s1342_) >>= fun w__64 : option ((word_width * bool * bool * mword 5 * mword 5)) =>
+ (if match w__64 with
+ | Some (size, aq, rl, rd, rs1) => true
+ | _ => false
+ end then
+ (_s1600_ _s1342_) >>= fun w__65 : option ((word_width * bool * bool * mword 5 * mword 5)) =>
(match w__65 with
- | Some ((size, aq, rl, rd, rs1)) =>
- returnm ((LOADRES
- ((aq, rl, rs1, size, rd)))
- : ast )
+ | Some (size, aq, rl, rd, rs1) =>
+ returnm (LOADRES (aq, rl, rs1, size, rd))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1101_ _s825_) >>= fun w__68 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__68 with
- | Some ((size, aq, rl, rd, rs1, rs2)) =>
- true
- | _ => false
- end)) then
- (_s1101_ _s825_) >>= fun w__69 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (_s1618_ _s1342_) >>= fun w__68 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (if match w__68 with
+ | Some (size, aq, rl, rd, rs1, rs2) =>
+ true
+ | _ => false
+ end then
+ (_s1618_ _s1342_) >>= fun w__69 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
(match w__69 with
- | Some ((size, aq, rl, rd, rs1, rs2)) =>
- returnm ((STORECON
- ((aq, rl, rs2, rs1, size, rd)))
- : ast )
+ | Some (size, aq, rl, rd, rs1, rs2) =>
+ returnm (STORECON
+ (aq, rl, rs2, rs1, size, rd))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1123_ _s825_) >>= fun w__72 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__72 with
- | Some
- ((op, width, aq, rl, rd, rs1, rs2)) =>
- true
- | _ => false
- end)) then
- (_s1123_ _s825_) >>= fun w__73 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (_s1640_ _s1342_) >>= fun w__72 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (if match w__72 with
+ | Some
+ (op, width, aq, rl, rd, rs2, rs1) =>
+ true
+ | _ => false
+ end then
+ (_s1640_ _s1342_) >>= fun w__73 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
(match w__73 with
| Some
- ((op, width, aq, rl, rd, rs1, rs2)) =>
- returnm ((AMO
- ((op, aq, rl, rs2, rs1, width, rd)))
- : ast )
+ (op, width, aq, rl, rd, rs2, rs1) =>
+ returnm (AMO
+ (op, aq, rl, rs2, rs1, width, rd))
| _ => exit tt : M (ast)
end)
: M (ast)
- else if ((generic_eq _s825_ "c.nop")) then
- returnm ((C_NOP
- (tt))
- : ast )
+ else if generic_eq _s1342_ "c.nop" then
+ returnm (C_NOP tt)
else
- (_s1147_ _s825_) >>= fun w__76 : option ((mword 3 * mword 8)) =>
- (if ((match w__76 with
- | Some ((rdc, nzimm)) =>
- neq_vec nzimm
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0]
- : mword 8)
- | _ => false
- end)) then
- (_s1147_ _s825_) >>= fun w__77 : option ((mword 3 * mword 8)) =>
+ (_s1666_ _s1342_) >>= fun w__76 : option ((mword 3 * mword 8)) =>
+ (if match w__76 with
+ | Some (rdc, nzimm) =>
+ neq_vec nzimm (Ox"00" : mword 8)
+ | _ => false
+ end then
+ (_s1666_ _s1342_) >>= fun w__77 : option ((mword 3 * mword 8)) =>
(match w__77 with
- | Some ((rdc, nzimm)) =>
- returnm ((C_ADDI4SPN
- ((rdc, nzimm)))
- : ast )
+ | Some (rdc, nzimm) =>
+ returnm (C_ADDI4SPN (rdc, nzimm))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1159_ _s825_) >>= fun w__80 : option ((mword 3 * mword 3 * mword 5)) =>
- (if ((match w__80 with
- | Some ((rdc, rsc, uimm)) =>
- true
- | _ => false
- end)) then
- (_s1159_ _s825_) >>= fun w__81 : option ((mword 3 * mword 3 * mword 5)) =>
+ (_s1678_ _s1342_) >>= fun w__80 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if match w__80 with
+ | Some (rdc, rsc, uimm) => true
+ | _ => false
+ end then
+ (_s1678_ _s1342_) >>= fun w__81 : option ((mword 3 * mword 3 * mword 5)) =>
(match w__81 with
- | Some ((rdc, rsc, uimm)) =>
- returnm ((C_LW
- ((uimm, rsc, rdc)))
- : ast )
+ | Some (rdc, rsc, uimm) =>
+ returnm (C_LW (uimm, rsc, rdc))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1175_ _s825_) >>= fun w__84 : option ((mword 3 * mword 3 * mword 5)) =>
- (if ((match w__84 with
- | Some ((rdc, rsc, uimm)) =>
- Z.eqb 64 64
- | _ => false
- end)) then
- (_s1175_ _s825_) >>= fun w__85 : option ((mword 3 * mword 3 * mword 5)) =>
+ (_s1694_ _s1342_) >>= fun w__84 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if match w__84 with
+ | Some (rdc, rsc, uimm) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s1694_ _s1342_) >>= fun w__85 : option ((mword 3 * mword 3 * mword 5)) =>
(match w__85 with
- | Some ((rdc, rsc, uimm)) =>
- returnm ((C_LD
- ((uimm, rsc, rdc)))
- : ast )
+ | Some (rdc, rsc, uimm) =>
+ returnm (C_LD
+ (uimm, rsc, rdc))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1191_ _s825_) >>= fun w__88 : option ((mword 3 * mword 3 * mword 5)) =>
- (if ((match w__88 with
- | Some
- ((rsc1, rsc2, uimm)) =>
- true
- | _ => false
- end)) then
- (_s1191_ _s825_) >>= fun w__89 : option ((mword 3 * mword 3 * mword 5)) =>
+ (_s1710_ _s1342_) >>= fun w__88 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if match w__88 with
+ | Some (rsc1, rsc2, uimm) =>
+ true
+ | _ => false
+ end then
+ (_s1710_ _s1342_) >>= fun w__89 : option ((mword 3 * mword 3 * mword 5)) =>
(match w__89 with
- | Some ((rsc1, rsc2, uimm)) =>
- returnm ((C_SW
- ((uimm, rsc1, rsc2)))
- : ast )
+ | Some (rsc1, rsc2, uimm) =>
+ returnm (C_SW
+ (uimm, rsc1, rsc2))
| _ => exit tt : M (ast)
end)
: M (ast)
else
- (_s1207_ _s825_) >>= fun w__92 : option ((mword 3 * mword 3 * mword 5)) =>
- (if ((match w__92 with
- | Some
- ((rsc1, rsc2, uimm)) =>
- Z.eqb 64 64
- | _ => false
- end)) then
- (_s1207_ _s825_) >>= fun w__93 : option ((mword 3 * mword 3 * mword 5)) =>
+ (_s1726_ _s1342_) >>= fun w__92 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if match w__92 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s1726_ _s1342_) >>= fun w__93 : option ((mword 3 * mword 3 * mword 5)) =>
(match w__93 with
| Some
- ((rsc1, rsc2, uimm)) =>
- returnm ((C_SD
- ((uimm, rsc1, rsc2)))
- : ast )
+ (rsc1, rsc2, uimm) =>
+ returnm (C_SD
+ (uimm, rsc1, rsc2))
| _ =>
exit tt : M (ast)
end)
: M (ast)
else
- (_s1223_ _s825_) >>= fun w__96 : option ((mword 5 * mword 6)) =>
- (if ((match w__96 with
- | Some ((rsd, nzi)) =>
- andb
- (neq_vec nzi
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ => false
- end)) then
- (_s1223_ _s825_) >>= fun w__97 : option ((mword 5 * mword 6)) =>
+ (_s1742_ _s1342_) >>= fun w__96 : option ((mword 5 * mword 6)) =>
+ (if match w__96 with
+ | Some (rsd, nzi) =>
+ andb
+ (neq_vec nzi
+ ('b"000000"
+ : mword 6))
+ (neq_vec rsd
+ zreg)
+ | _ => false
+ end then
+ (_s1742_ _s1342_) >>= fun w__97 : option ((mword 5 * mword 6)) =>
(match w__97 with
- | Some ((rsd, nzi)) =>
- returnm ((C_ADDI
- ((nzi, rsd)))
- : ast )
+ | Some (rsd, nzi) =>
+ returnm (C_ADDI
+ (nzi, rsd))
| _ =>
exit tt : M (ast)
end)
: M (ast)
else
- (_s1235_ _s825_) >>= fun w__100 : option (mword 11) =>
- (if ((match w__100 with
- | Some (imm) =>
- Z.eqb 64 32
- | _ => false
- end)) then
- (_s1235_ _s825_) >>= fun w__101 : option (mword 11) =>
+ (_s1754_ _s1342_) >>= fun w__100 : option (mword 11) =>
+ (if match w__100 with
+ | Some imm =>
+ Z.eqb 64 32
+ | _ => false
+ end then
+ (_s1754_ _s1342_) >>= fun w__101 : option (mword 11) =>
(match w__101 with
- | Some (imm) =>
- returnm ((C_JAL
- (imm))
- : ast )
+ | Some imm =>
+ returnm (C_JAL
+ imm)
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1243_ _s825_) >>= fun w__104 : option ((mword 5 * mword 6)) =>
- (if ((match w__104 with
- | Some
- ((rsd, imm)) =>
- Z.eqb 64
- 64
- | _ => false
- end)) then
- (_s1243_ _s825_) >>= fun w__105 : option ((mword 5 * mword 6)) =>
+ (_s1762_ _s1342_) >>= fun w__104 : option ((mword 5 * mword 6)) =>
+ (if match w__104 with
+ | Some
+ (rsd, imm) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s1762_ _s1342_) >>= fun w__105 : option ((mword 5 * mword 6)) =>
(match w__105 with
| Some
- ((rsd, imm)) =>
- returnm ((C_ADDIW
- ((imm, rsd)))
- : ast )
+ (rsd, imm) =>
+ returnm (C_ADDIW
+ (imm, rsd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1255_ _s825_) >>= fun w__108 : option ((mword 5 * mword 6)) =>
- (if ((match w__108 with
- | Some
- ((rd, imm)) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end)) then
- (_s1255_
- _s825_) >>= fun w__109 : option ((mword 5 * mword 6)) =>
+ (_s1774_ _s1342_) >>= fun w__108 : option ((mword 5 * mword 6)) =>
+ (if match w__108 with
+ | Some
+ (rd, imm) =>
+ neq_vec
+ rd zreg
+ | _ => false
+ end then
+ (_s1774_
+ _s1342_) >>= fun w__109 : option ((mword 5 * mword 6)) =>
(match w__109 with
| Some
- ((rd, imm)) =>
- returnm ((C_LI
- ((imm, rd)))
- : ast )
+ (rd, imm) =>
+ returnm (C_LI
+ (imm, rd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1267_
- _s825_) >>= fun w__112 : option (mword 6) =>
- (if ((match w__112 with
- | Some
- (imm) =>
- neq_vec
- imm
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
- then
- (_s1267_
- _s825_) >>= fun w__113 : option (mword 6) =>
+ (_s1786_
+ _s1342_) >>= fun w__112 : option (mword 6) =>
+ (if match w__112 with
+ | Some
+ imm =>
+ neq_vec
+ imm
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end then
+ (_s1786_
+ _s1342_) >>= fun w__113 : option (mword 6) =>
(match w__113 with
| Some
- (imm) =>
- returnm ((C_ADDI16SP
- (imm))
- : ast )
+ imm =>
+ returnm (C_ADDI16SP
+ imm)
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1275_
- _s825_) >>= fun w__116 : option ((mword 5 * mword 6)) =>
- (if ((match w__116 with
- | Some
- ((rd, imm)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- sp))))
- (neq_vec
- imm
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)))
- : bool)
- | _ =>
- false
- end))
+ (_s1794_
+ _s1342_) >>= fun w__116 : option ((mword 5 * mword 6)) =>
+ (if match w__116 with
+ | Some
+ (rd, imm) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (andb
+ (neq_vec
+ rd
+ sp)
+ (neq_vec
+ imm
+ ('b"000000"
+ : mword 6)))
+ | _ =>
+ false
+ end
then
- (_s1275_
- _s825_) >>= fun w__117 : option ((mword 5 * mword 6)) =>
+ (_s1794_
+ _s1342_) >>= fun w__117 : option ((mword 5 * mword 6)) =>
(match w__117 with
| Some
- ((rd, imm)) =>
- returnm ((C_LUI
- ((imm, rd)))
- : ast )
+ (rd, imm) =>
+ returnm (C_LUI
+ (imm, rd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1287_
- _s825_) >>= fun w__120 : option ((mword 3 * mword 6)) =>
- (if ((match w__120 with
- | Some
- ((rsd, shamt)) =>
- neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
+ (_s1806_
+ _s1342_) >>= fun w__120 : option ((mword 3 * mword 6)) =>
+ (if match w__120 with
+ | Some
+ (rsd, shamt) =>
+ neq_vec
+ shamt
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end
then
- (_s1287_
- _s825_) >>= fun w__121 : option ((mword 3 * mword 6)) =>
+ (_s1806_
+ _s1342_) >>= fun w__121 : option ((mword 3 * mword 6)) =>
(match w__121 with
| Some
- ((rsd, shamt)) =>
- returnm ((C_SRLI
- ((shamt, rsd)))
- : ast )
+ (rsd, shamt) =>
+ returnm (C_SRLI
+ (shamt, rsd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1299_
- _s825_) >>= fun w__124 : option ((mword 3 * mword 6)) =>
- (if ((match w__124 with
- | Some
- ((rsd, shamt)) =>
- neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
+ (_s1818_
+ _s1342_) >>= fun w__124 : option ((mword 3 * mword 6)) =>
+ (if match w__124 with
+ | Some
+ (rsd, shamt) =>
+ neq_vec
+ shamt
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end
then
- (_s1299_
- _s825_) >>= fun w__125 : option ((mword 3 * mword 6)) =>
+ (_s1818_
+ _s1342_) >>= fun w__125 : option ((mword 3 * mword 6)) =>
(match w__125 with
| Some
- ((rsd, shamt)) =>
- returnm ((C_SRAI
- ((shamt, rsd)))
- : ast )
+ (rsd, shamt) =>
+ returnm (C_SRAI
+ (shamt, rsd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1311_
- _s825_) >>= fun w__128 : option ((mword 3 * mword 6)) =>
+ (_s1830_
+ _s1342_) >>= fun w__128 : option ((mword 3 * mword 6)) =>
(if
- ((match w__128 with
- | Some
- ((rsd, imm)) =>
- true
- | _ =>
- false
- end))
+ match w__128 with
+ | Some
+ (rsd, imm) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1311_
- _s825_) >>= fun w__129 : option ((mword 3 * mword 6)) =>
+ (_s1830_
+ _s1342_) >>= fun w__129 : option ((mword 3 * mword 6)) =>
(match w__129 with
| Some
- ((rsd, imm)) =>
- returnm ((C_ANDI
- ((imm, rsd)))
- : ast )
+ (rsd, imm) =>
+ returnm (C_ANDI
+ (imm, rsd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1323_
- _s825_) >>= fun w__132 : option ((mword 3 * mword 3)) =>
+ (_s1842_
+ _s1342_) >>= fun w__132 : option ((mword 3 * mword 3)) =>
(if
- ((match w__132 with
- | Some
- ((rsd, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__132 with
+ | Some
+ (rsd, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1323_
- _s825_) >>= fun w__133 : option ((mword 3 * mword 3)) =>
+ (_s1842_
+ _s1342_) >>= fun w__133 : option ((mword 3 * mword 3)) =>
(match w__133 with
| Some
- ((rsd, rs2)) =>
- returnm ((C_SUB
- ((rsd, rs2)))
- : ast )
+ (rsd, rs2) =>
+ returnm (C_SUB
+ (rsd, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1335_
- _s825_) >>= fun w__136 : option ((mword 3 * mword 3)) =>
+ (_s1854_
+ _s1342_) >>= fun w__136 : option ((mword 3 * mword 3)) =>
(if
- ((match w__136 with
- | Some
- ((rsd, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__136 with
+ | Some
+ (rsd, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1335_
- _s825_) >>= fun w__137 : option ((mword 3 * mword 3)) =>
+ (_s1854_
+ _s1342_) >>= fun w__137 : option ((mword 3 * mword 3)) =>
(match w__137 with
| Some
- ((rsd, rs2)) =>
- returnm ((C_XOR
- ((rsd, rs2)))
- : ast )
+ (rsd, rs2) =>
+ returnm (C_XOR
+ (rsd, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1347_
- _s825_) >>= fun w__140 : option ((mword 3 * mword 3)) =>
+ (_s1866_
+ _s1342_) >>= fun w__140 : option ((mword 3 * mword 3)) =>
(if
- ((match w__140 with
- | Some
- ((rsd, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__140 with
+ | Some
+ (rsd, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1347_
- _s825_) >>= fun w__141 : option ((mword 3 * mword 3)) =>
+ (_s1866_
+ _s1342_) >>= fun w__141 : option ((mword 3 * mword 3)) =>
(match w__141 with
| Some
- ((rsd, rs2)) =>
- returnm ((C_OR
- ((rsd, rs2)))
- : ast )
+ (rsd, rs2) =>
+ returnm (C_OR
+ (rsd, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1359_
- _s825_) >>= fun w__144 : option ((mword 3 * mword 3)) =>
+ (_s1878_
+ _s1342_) >>= fun w__144 : option ((mword 3 * mword 3)) =>
(if
- ((match w__144 with
- | Some
- ((rsd, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__144 with
+ | Some
+ (rsd, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1359_
- _s825_) >>= fun w__145 : option ((mword 3 * mword 3)) =>
+ (_s1878_
+ _s1342_) >>= fun w__145 : option ((mword 3 * mword 3)) =>
(match w__145 with
| Some
- ((rsd, rs2)) =>
- returnm ((C_AND
- ((rsd, rs2)))
- : ast )
+ (rsd, rs2) =>
+ returnm (C_AND
+ (rsd, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1371_
- _s825_) >>= fun w__148 : option ((mword 3 * mword 3)) =>
+ (_s1890_
+ _s1342_) >>= fun w__148 : option ((mword 3 * mword 3)) =>
(if
- ((match w__148 with
- | Some
- ((rsd, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__148 with
+ | Some
+ (rsd, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s1371_
- _s825_) >>= fun w__149 : option ((mword 3 * mword 3)) =>
+ (_s1890_
+ _s1342_) >>= fun w__149 : option ((mword 3 * mword 3)) =>
(match w__149 with
| Some
- ((rsd, rs2)) =>
- returnm ((C_SUBW
- ((rsd, rs2)))
- : ast )
+ (rsd, rs2) =>
+ returnm (C_SUBW
+ (rsd, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1383_
- _s825_) >>= fun w__152 : option ((mword 3 * mword 3)) =>
+ (_s1902_
+ _s1342_) >>= fun w__152 : option ((mword 3 * mword 3)) =>
(if
- ((match w__152 with
- | Some
- ((rsd, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__152 with
+ | Some
+ (rsd, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s1383_
- _s825_) >>= fun w__153 : option ((mword 3 * mword 3)) =>
+ (_s1902_
+ _s1342_) >>= fun w__153 : option ((mword 3 * mword 3)) =>
(match w__153 with
| Some
- ((rsd, rs2)) =>
- returnm ((C_ADDW
- ((rsd, rs2)))
- : ast )
+ (rsd, rs2) =>
+ returnm (C_ADDW
+ (rsd, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1395_
- _s825_) >>= fun w__156 : option (mword 11) =>
+ (_s1914_
+ _s1342_) >>= fun w__156 : option (mword 11) =>
(if
- ((match w__156 with
- | Some
- (imm) =>
- true
- | _ =>
- false
- end))
+ match w__156 with
+ | Some
+ imm =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1395_
- _s825_) >>= fun w__157 : option (mword 11) =>
+ (_s1914_
+ _s1342_) >>= fun w__157 : option (mword 11) =>
(match w__157 with
| Some
- (imm) =>
- returnm ((C_J
- (imm))
- : ast )
+ imm =>
+ returnm (C_J
+ imm)
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1403_
- _s825_) >>= fun w__160 : option ((mword 3 * mword 8)) =>
+ (_s1922_
+ _s1342_) >>= fun w__160 : option ((mword 3 * mword 8)) =>
(if
- ((match w__160 with
- | Some
- ((rs, imm)) =>
- true
- | _ =>
- false
- end))
+ match w__160 with
+ | Some
+ (rs, imm) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1403_
- _s825_) >>= fun w__161 : option ((mword 3 * mword 8)) =>
+ (_s1922_
+ _s1342_) >>= fun w__161 : option ((mword 3 * mword 8)) =>
(match w__161 with
| Some
- ((rs, imm)) =>
- returnm ((C_BEQZ
- ((imm, rs)))
- : ast )
+ (rs, imm) =>
+ returnm (C_BEQZ
+ (imm, rs))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1415_
- _s825_) >>= fun w__164 : option ((mword 3 * mword 8)) =>
+ (_s1934_
+ _s1342_) >>= fun w__164 : option ((mword 3 * mword 8)) =>
(if
- ((match w__164 with
- | Some
- ((rs, imm)) =>
- true
- | _ =>
- false
- end))
+ match w__164 with
+ | Some
+ (rs, imm) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1415_
- _s825_) >>= fun w__165 : option ((mword 3 * mword 8)) =>
+ (_s1934_
+ _s1342_) >>= fun w__165 : option ((mword 3 * mword 8)) =>
(match w__165 with
| Some
- ((rs, imm)) =>
- returnm ((C_BNEZ
- ((imm, rs)))
- : ast )
+ (rs, imm) =>
+ returnm (C_BNEZ
+ (imm, rs))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1427_
- _s825_) >>= fun w__168 : option ((mword 5 * mword 6)) =>
+ (_s1946_
+ _s1342_) >>= fun w__168 : option ((mword 5 * mword 6)) =>
(if
- ((match w__168 with
- | Some
- ((rsd, shamt)) =>
- andb
- (neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__168 with
+ | Some
+ (rsd, shamt) =>
+ andb
+ (neq_vec
+ shamt
+ ('b"000000"
+ : mword 6))
+ (neq_vec
+ rsd
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s1427_
- _s825_) >>= fun w__169 : option ((mword 5 * mword 6)) =>
+ (_s1946_
+ _s1342_) >>= fun w__169 : option ((mword 5 * mword 6)) =>
(match w__169 with
| Some
- ((rsd, shamt)) =>
- returnm ((C_SLLI
- ((shamt, rsd)))
- : ast )
+ (rsd, shamt) =>
+ returnm (C_SLLI
+ (shamt, rsd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1439_
- _s825_) >>= fun w__172 : option ((mword 5 * mword 6)) =>
+ (_s1958_
+ _s1342_) >>= fun w__172 : option ((mword 5 * mword 6)) =>
(if
- ((match w__172 with
- | Some
- ((rd, uimm)) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__172 with
+ | Some
+ (rd, uimm) =>
+ neq_vec
+ rd
+ zreg
+ | _ =>
+ false
+ end
then
- (_s1439_
- _s825_) >>= fun w__173 : option ((mword 5 * mword 6)) =>
+ (_s1958_
+ _s1342_) >>= fun w__173 : option ((mword 5 * mword 6)) =>
(match w__173 with
| Some
- ((rd, uimm)) =>
- returnm ((C_LWSP
- ((uimm, rd)))
- : ast )
+ (rd, uimm) =>
+ returnm (C_LWSP
+ (uimm, rd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1451_
- _s825_) >>= fun w__176 : option ((mword 5 * mword 6)) =>
+ (_s1970_
+ _s1342_) >>= fun w__176 : option ((mword 5 * mword 6)) =>
(if
- ((match w__176 with
- | Some
- ((rd, uimm)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- (Z.eqb
- 64
- 64)
- | _ =>
- false
- end))
+ match w__176 with
+ | Some
+ (rd, uimm) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
then
- (_s1451_
- _s825_) >>= fun w__177 : option ((mword 5 * mword 6)) =>
+ (_s1970_
+ _s1342_) >>= fun w__177 : option ((mword 5 * mword 6)) =>
(match w__177 with
| Some
- ((rd, uimm)) =>
- returnm ((C_LDSP
- ((uimm, rd)))
- : ast )
+ (rd, uimm) =>
+ returnm (C_LDSP
+ (uimm, rd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1463_
- _s825_) >>= fun w__180 : option ((mword 5 * mword 6)) =>
+ (_s1982_
+ _s1342_) >>= fun w__180 : option ((mword 5 * mword 6)) =>
(if
- ((match w__180 with
- | Some
- ((rd, uimm)) =>
- true
- | _ =>
- false
- end))
+ match w__180 with
+ | Some
+ (rd, uimm) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1463_
- _s825_) >>= fun w__181 : option ((mword 5 * mword 6)) =>
+ (_s1982_
+ _s1342_) >>= fun w__181 : option ((mword 5 * mword 6)) =>
(match w__181 with
| Some
- ((rd, uimm)) =>
- returnm ((C_SWSP
- ((uimm, rd)))
- : ast )
+ (rd, uimm) =>
+ returnm (C_SWSP
+ (uimm, rd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1475_
- _s825_) >>= fun w__184 : option ((mword 5 * mword 6)) =>
+ (_s1994_
+ _s1342_) >>= fun w__184 : option ((mword 5 * mword 6)) =>
(if
- ((match w__184 with
- | Some
- ((rs2, uimm)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__184 with
+ | Some
+ (rs2, uimm) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s1475_
- _s825_) >>= fun w__185 : option ((mword 5 * mword 6)) =>
+ (_s1994_
+ _s1342_) >>= fun w__185 : option ((mword 5 * mword 6)) =>
(match w__185 with
| Some
- ((rs2, uimm)) =>
- returnm ((C_SDSP
- ((uimm, rs2)))
- : ast )
+ (rs2, uimm) =>
+ returnm (C_SDSP
+ (uimm, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1487_
- _s825_) >>= fun w__188 : option (mword 5) =>
+ (_s2006_
+ _s1342_) >>= fun w__188 : option (mword 5) =>
(if
- ((match w__188 with
- | Some
- (rs1) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs1))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__188 with
+ | Some
+ rs1 =>
+ neq_vec
+ rs1
+ zreg
+ | _ =>
+ false
+ end
then
- (_s1487_
- _s825_) >>= fun w__189 : option (mword 5) =>
+ (_s2006_
+ _s1342_) >>= fun w__189 : option (mword 5) =>
(match w__189 with
| Some
- (rs1) =>
- returnm ((C_JR
- (rs1))
- : ast )
+ rs1 =>
+ returnm (C_JR
+ rs1)
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1495_
- _s825_) >>= fun w__192 : option (mword 5) =>
+ (_s2014_
+ _s1342_) >>= fun w__192 : option (mword 5) =>
(if
- ((match w__192 with
- | Some
- (rs1) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs1))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__192 with
+ | Some
+ rs1 =>
+ neq_vec
+ rs1
+ zreg
+ | _ =>
+ false
+ end
then
- (_s1495_
- _s825_) >>= fun w__193 : option (mword 5) =>
+ (_s2014_
+ _s1342_) >>= fun w__193 : option (mword 5) =>
(match w__193 with
| Some
- (rs1) =>
- returnm ((C_JALR
- (rs1))
- : ast )
+ rs1 =>
+ returnm (C_JALR
+ rs1)
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1503_
- _s825_) >>= fun w__196 : option ((mword 5 * mword 5)) =>
+ (_s2022_
+ _s1342_) >>= fun w__196 : option ((mword 5 * mword 5)) =>
(if
- ((match w__196 with
- | Some
- ((rd, rs2)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs2))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__196 with
+ | Some
+ (rd, rs2) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (neq_vec
+ rs2
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s1503_
- _s825_) >>= fun w__197 : option ((mword 5 * mword 5)) =>
+ (_s2022_
+ _s1342_) >>= fun w__197 : option ((mword 5 * mword 5)) =>
(match w__197 with
| Some
- ((rd, rs2)) =>
- returnm ((C_MV
- ((rd, rs2)))
- : ast )
+ (rd, rs2) =>
+ returnm (C_MV
+ (rd, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else if
- ((generic_eq
- _s825_
- "c.ebreak"))
+ generic_eq
+ _s1342_
+ "c.ebreak"
then
- returnm ((C_EBREAK
- (tt))
- : ast )
+ returnm (C_EBREAK
+ tt)
else
- (_s1515_
- _s825_) >>= fun w__200 : option ((mword 5 * mword 5)) =>
+ (_s2034_
+ _s1342_) >>= fun w__200 : option ((mword 5 * mword 5)) =>
(if
- ((match w__200 with
- | Some
- ((rsd, rs2)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs2))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__200 with
+ | Some
+ (rsd, rs2) =>
+ andb
+ (neq_vec
+ rsd
+ zreg)
+ (neq_vec
+ rs2
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s1515_
- _s825_) >>= fun w__201 : option ((mword 5 * mword 5)) =>
+ (_s2034_
+ _s1342_) >>= fun w__201 : option ((mword 5 * mword 5)) =>
(match w__201 with
| Some
- ((rsd, rs2)) =>
- returnm ((C_ADD
- ((rsd, rs2)))
- : ast )
+ (rsd, rs2) =>
+ returnm (C_ADD
+ (rsd, rs2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1527_
- _s825_) >>= fun w__204 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2046_
+ _s1342_) >>= fun w__204 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__204 with
- | Some
- ((high, signed1, signed2, rd, rs1, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__204 with
+ | Some
+ (high, signed1, signed2, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1527_
- _s825_) >>= fun w__205 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2046_
+ _s1342_) >>= fun w__205 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)) =>
(match w__205 with
| Some
- ((high, signed1, signed2, rd, rs1, rs2)) =>
- returnm ((MUL
- ((rs2, rs1, rd, high, signed1, signed2)))
- : ast )
+ (high, signed1, signed2, rd, rs1, rs2) =>
+ returnm (MUL
+ (rs2, rs1, rd, high, signed1, signed2))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1544_
- _s825_) >>= fun w__208 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2063_
+ _s1342_) >>= fun w__208 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__208 with
- | Some
- ((s, rd, rs1, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__208 with
+ | Some
+ (s, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1544_
- _s825_) >>= fun w__209 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2063_
+ _s1342_) >>= fun w__209 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(match w__209 with
| Some
- ((s, rd, rs1, rs2)) =>
- returnm ((DIV
- ((rs2, rs1, rd, s)))
- : ast )
+ (s, rd, rs1, rs2) =>
+ returnm (DIV
+ (rs2, rs1, rd, s))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1562_
- _s825_) >>= fun w__212 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2081_
+ _s1342_) >>= fun w__212 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__212 with
- | Some
- ((s, rd, rs1, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__212 with
+ | Some
+ (s, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1562_
- _s825_) >>= fun w__213 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2081_
+ _s1342_) >>= fun w__213 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(match w__213 with
| Some
- ((s, rd, rs1, rs2)) =>
- returnm ((REM
- ((rs2, rs1, rd, s)))
- : ast )
+ (s, rd, rs1, rs2) =>
+ returnm (REM
+ (rs2, rs1, rd, s))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1580_
- _s825_) >>= fun w__216 : option ((mword 5 * mword 5 * mword 5)) =>
+ (_s2099_
+ _s1342_) >>= fun w__216 : option ((mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__216 with
- | Some
- ((rd, rs1, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__216 with
+ | Some
+ (rd, rs1, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s1580_
- _s825_) >>= fun w__217 : option ((mword 5 * mword 5 * mword 5)) =>
+ (_s2099_
+ _s1342_) >>= fun w__217 : option ((mword 5 * mword 5 * mword 5)) =>
(match w__217 with
| Some
- ((rd, rs1, rs2)) =>
- returnm ((MULW
- ((rs2, rs1, rd)))
- : ast )
+ (rd, rs1, rs2) =>
+ returnm (MULW
+ (rs2, rs1, rd))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1596_
- _s825_) >>= fun w__220 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2115_
+ _s1342_) >>= fun w__220 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__220 with
- | Some
- ((s, rd, rs1, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__220 with
+ | Some
+ (s, rd, rs1, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s1596_
- _s825_) >>= fun w__221 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2115_
+ _s1342_) >>= fun w__221 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(match w__221 with
| Some
- ((s, rd, rs1, rs2)) =>
- returnm ((DIVW
- ((rs2, rs1, rd, s)))
- : ast )
+ (s, rd, rs1, rs2) =>
+ returnm (DIVW
+ (rs2, rs1, rd, s))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1615_
- _s825_) >>= fun w__224 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2134_
+ _s1342_) >>= fun w__224 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__224 with
- | Some
- ((s, rd, rs1, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__224 with
+ | Some
+ (s, rd, rs1, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s1615_
- _s825_) >>= fun w__225 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s2134_
+ _s1342_) >>= fun w__225 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(match w__225 with
| Some
- ((s, rd, rs1, rs2)) =>
- returnm ((REMW
- ((rs2, rs1, rd, s)))
- : ast )
+ (s, rd, rs1, rs2) =>
+ returnm (REMW
+ (rs2, rs1, rd, s))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1634_
- _s825_) >>= fun w__228 : option ((csrop * mword 5 * mword 5 * mword 12)) =>
+ (_s2153_
+ _s1342_) >>= fun w__228 : option ((csrop * mword 5 * mword 12 * mword 5)) =>
(if
- ((match w__228 with
- | Some
- ((op, rd, rs1, csr)) =>
- true
- | _ =>
- false
- end))
+ match w__228 with
+ | Some
+ (op, rd, csr, rs1) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1634_
- _s825_) >>= fun w__229 : option ((csrop * mword 5 * mword 5 * mword 12)) =>
+ (_s2153_
+ _s1342_) >>= fun w__229 : option ((csrop * mword 5 * mword 12 * mword 5)) =>
(match w__229 with
| Some
- ((op, rd, rs1, csr)) =>
- returnm ((CSR
- ((csr, rs1, rd, true, op)))
- : ast )
+ (op, rd, csr, rs1) =>
+ returnm (CSR
+ (csr, rs1, rd, true, op))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1652_
- _s825_) >>= fun w__232 : option ((csrop * mword 5 * mword 5 * mword 12)) =>
+ (_s2171_
+ _s1342_) >>= fun w__232 : option ((csrop * mword 5 * mword 12 * mword 5)) =>
(if
- ((match w__232 with
- | Some
- ((op, rd, rs1, csr)) =>
- true
- | _ =>
- false
- end))
+ match w__232 with
+ | Some
+ (op, rd, csr, rs1) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1652_
- _s825_) >>= fun w__233 : option ((csrop * mword 5 * mword 5 * mword 12)) =>
+ (_s2171_
+ _s1342_) >>= fun w__233 : option ((csrop * mword 5 * mword 12 * mword 5)) =>
(match w__233 with
| Some
- ((op, rd, rs1, csr)) =>
- returnm ((CSR
- ((csr, rs1, rd, false, op)))
- : ast )
+ (op, rd, csr, rs1) =>
+ returnm (CSR
+ (csr, rs1, rd, false, op))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else if
- ((generic_eq
- _s825_
- "uret"))
+ generic_eq
+ _s1342_
+ "uret"
then
- returnm ((URET
- (tt))
- : ast )
+ returnm (URET
+ tt)
else
- (_s1669_
- _s825_) >>= fun w__236 : option (mword 32) =>
+ (_s2188_
+ _s1342_) >>= fun w__236 : option ((word_width * mword 5 * mword 12 * mword 5)) =>
(if
- ((match w__236 with
- | Some
- (s) =>
- true
- | _ =>
- false
- end))
+ match w__236 with
+ | Some
+ (width, rd, imm, rs1) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1669_
- _s825_) >>= fun w__237 : option (mword 32) =>
+ (_s2188_
+ _s1342_) >>= fun w__237 : option ((word_width * mword 5 * mword 12 * mword 5)) =>
(match w__237 with
| Some
- (s) =>
- returnm ((ILLEGAL
- (s))
- : ast )
+ (width, rd, imm, rs1) =>
+ returnm (LOAD_FP
+ (imm, rs1, rd, width))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- (_s1677_
- _s825_) >>= fun w__240 : option (mword 16) =>
+ (_s2212_
+ _s1342_) >>= fun w__240 : option ((word_width * mword 5 * mword 12 * mword 5)) =>
(if
- ((match w__240 with
- | Some
- (s) =>
- true
- | _ =>
- false
- end))
+ match w__240 with
+ | Some
+ (width, rs2, imm, rs1) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s1677_
- _s825_) >>= fun w__241 : option (mword 16) =>
+ (_s2212_
+ _s1342_) >>= fun w__241 : option ((word_width * mword 5 * mword 12 * mword 5)) =>
(match w__241 with
| Some
- (s) =>
- returnm ((C_ILLEGAL
- (s))
- : ast )
+ (width, rs2, imm, rs1) =>
+ returnm (STORE_FP
+ (imm, rs2, rs1, width))
| _ =>
exit tt
: M (ast)
end)
: M (ast)
else
- assert_exp' false "Pattern match failure at unknown location" >>= fun _ =>
- exit tt)
+ (_s2236_
+ _s1342_) >>= fun w__244 : option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__244 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2236_
+ _s1342_) >>= fun w__245 : option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__245 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm) =>
+ returnm (F_MADD_TYPE_S
+ (rs3, rs2, rs1, rm, rd, op))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2261_
+ _s1342_) >>= fun w__248 : option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__248 with
+ | Some
+ (op, rd, rs1, rs2, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2261_
+ _s1342_) >>= fun w__249 : option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__249 with
+ | Some
+ (op, rd, rs1, rs2, rm) =>
+ returnm (F_BIN_RM_TYPE_S
+ (rs2, rs1, rm, rd, op))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2282_
+ _s1342_) >>= fun w__252 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__252 with
+ | Some
+ (FSQRT_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2282_
+ _s1342_) >>= fun w__253 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__253 with
+ | Some
+ (FSQRT_S, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FSQRT_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2299_
+ _s1342_) >>= fun w__256 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__256 with
+ | Some
+ (FCVT_W_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2299_
+ _s1342_) >>= fun w__257 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__257 with
+ | Some
+ (FCVT_W_S, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_W_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2316_
+ _s1342_) >>= fun w__260 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__260 with
+ | Some
+ (FCVT_WU_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2316_
+ _s1342_) >>= fun w__261 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__261 with
+ | Some
+ (FCVT_WU_S, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_WU_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2333_
+ _s1342_) >>= fun w__264 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__264 with
+ | Some
+ (FCVT_S_W, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2333_
+ _s1342_) >>= fun w__265 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__265 with
+ | Some
+ (FCVT_S_W, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_W))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2350_
+ _s1342_) >>= fun w__268 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__268 with
+ | Some
+ (FCVT_S_WU, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2350_
+ _s1342_) >>= fun w__269 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__269 with
+ | Some
+ (FCVT_S_WU, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_WU))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2367_
+ _s1342_) >>= fun w__272 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__272 with
+ | Some
+ (FCVT_L_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2367_
+ _s1342_) >>= fun w__273 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__273 with
+ | Some
+ (FCVT_L_S, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_L_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2384_
+ _s1342_) >>= fun w__276 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__276 with
+ | Some
+ (FCVT_LU_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2384_
+ _s1342_) >>= fun w__277 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__277 with
+ | Some
+ (FCVT_LU_S, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_LU_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2401_
+ _s1342_) >>= fun w__280 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__280 with
+ | Some
+ (FCVT_S_L, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2401_
+ _s1342_) >>= fun w__281 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__281 with
+ | Some
+ (FCVT_S_L, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_L))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2418_
+ _s1342_) >>= fun w__284 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__284 with
+ | Some
+ (FCVT_S_LU, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2418_
+ _s1342_) >>= fun w__285 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__285 with
+ | Some
+ (FCVT_S_LU, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_LU))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2435_
+ _s1342_) >>= fun w__288 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__288 with
+ | Some
+ (FSGNJ_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2435_
+ _s1342_) >>= fun w__289 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__289 with
+ | Some
+ (FSGNJ_S, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJ_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2452_
+ _s1342_) >>= fun w__292 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__292 with
+ | Some
+ (FSGNJN_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2452_
+ _s1342_) >>= fun w__293 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__293 with
+ | Some
+ (FSGNJN_S, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJN_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2469_
+ _s1342_) >>= fun w__296 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__296 with
+ | Some
+ (FSGNJX_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2469_
+ _s1342_) >>= fun w__297 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__297 with
+ | Some
+ (FSGNJX_S, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJX_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2486_
+ _s1342_) >>= fun w__300 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__300 with
+ | Some
+ (FMIN_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2486_
+ _s1342_) >>= fun w__301 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__301 with
+ | Some
+ (FMIN_S, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FMIN_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2503_
+ _s1342_) >>= fun w__304 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__304 with
+ | Some
+ (FMAX_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2503_
+ _s1342_) >>= fun w__305 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__305 with
+ | Some
+ (FMAX_S, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FMAX_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2520_
+ _s1342_) >>= fun w__308 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__308 with
+ | Some
+ (FEQ_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2520_
+ _s1342_) >>= fun w__309 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__309 with
+ | Some
+ (FEQ_S, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FEQ_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2537_
+ _s1342_) >>= fun w__312 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__312 with
+ | Some
+ (FLT_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2537_
+ _s1342_) >>= fun w__313 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__313 with
+ | Some
+ (FLT_S, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FLT_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2554_
+ _s1342_) >>= fun w__316 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__316 with
+ | Some
+ (FLE_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2554_
+ _s1342_) >>= fun w__317 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__317 with
+ | Some
+ (FLE_S, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_S
+ (rs2, rs1, rd, FLE_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2571_
+ _s1342_) >>= fun w__320 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (if
+ match w__320 with
+ | Some
+ (FMV_X_W, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2571_
+ _s1342_) >>= fun w__321 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (match w__321 with
+ | Some
+ (FMV_X_W, rd, rs1) =>
+ returnm (F_UN_TYPE_S
+ (rs1, rd, FMV_X_W))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2584_
+ _s1342_) >>= fun w__324 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (if
+ match w__324 with
+ | Some
+ (FMV_W_X, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2584_
+ _s1342_) >>= fun w__325 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (match w__325 with
+ | Some
+ (FMV_W_X, rd, rs1) =>
+ returnm (F_UN_TYPE_S
+ (rs1, rd, FMV_W_X))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2597_
+ _s1342_) >>= fun w__328 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (if
+ match w__328 with
+ | Some
+ (FCLASS_S, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2597_
+ _s1342_) >>= fun w__329 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (match w__329 with
+ | Some
+ (FCLASS_S, rd, rs1) =>
+ returnm (F_UN_TYPE_S
+ (rs1, rd, FCLASS_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2610_
+ _s1342_) >>= fun w__332 : option ((mword 5 * mword 6)) =>
+ (if
+ match w__332 with
+ | Some
+ (rd, imm) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s2610_
+ _s1342_) >>= fun w__333 : option ((mword 5 * mword 6)) =>
+ (match w__333 with
+ | Some
+ (rd, imm) =>
+ returnm (C_FLWSP
+ (imm, rd))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2622_
+ _s1342_) >>= fun w__336 : option ((mword 5 * mword 6)) =>
+ (if
+ match w__336 with
+ | Some
+ (rd, uimm) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s2622_
+ _s1342_) >>= fun w__337 : option ((mword 5 * mword 6)) =>
+ (match w__337 with
+ | Some
+ (rd, uimm) =>
+ returnm (C_FSWSP
+ (uimm, rd))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2634_
+ _s1342_) >>= fun w__340 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if
+ match w__340 with
+ | Some
+ (rdc, rsc, uimm) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s2634_
+ _s1342_) >>= fun w__341 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__341 with
+ | Some
+ (rdc, rsc, uimm) =>
+ returnm (C_FLW
+ (uimm, rsc, rdc))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2650_
+ _s1342_) >>= fun w__344 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if
+ match w__344 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s2650_
+ _s1342_) >>= fun w__345 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__345 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ returnm (C_FSW
+ (uimm, rsc1, rsc2))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2666_
+ _s1342_) >>= fun w__348 : option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__348 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2666_
+ _s1342_) >>= fun w__349 : option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__349 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm) =>
+ returnm (F_MADD_TYPE_D
+ (rs3, rs2, rs1, rm, rd, op))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2691_
+ _s1342_) >>= fun w__352 : option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__352 with
+ | Some
+ (op, rd, rs1, rs2, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2691_
+ _s1342_) >>= fun w__353 : option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__353 with
+ | Some
+ (op, rd, rs1, rs2, rm) =>
+ returnm (F_BIN_RM_TYPE_D
+ (rs2, rs1, rm, rd, op))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2712_
+ _s1342_) >>= fun w__356 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__356 with
+ | Some
+ (FSQRT_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2712_
+ _s1342_) >>= fun w__357 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__357 with
+ | Some
+ (FSQRT_D, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FSQRT_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2729_
+ _s1342_) >>= fun w__360 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__360 with
+ | Some
+ (FCVT_W_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2729_
+ _s1342_) >>= fun w__361 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__361 with
+ | Some
+ (FCVT_W_D, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_W_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2746_
+ _s1342_) >>= fun w__364 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__364 with
+ | Some
+ (FCVT_WU_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2746_
+ _s1342_) >>= fun w__365 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__365 with
+ | Some
+ (FCVT_WU_D, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_WU_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2763_
+ _s1342_) >>= fun w__368 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__368 with
+ | Some
+ (FCVT_D_W, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2763_
+ _s1342_) >>= fun w__369 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__369 with
+ | Some
+ (FCVT_D_W, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_W))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2780_
+ _s1342_) >>= fun w__372 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__372 with
+ | Some
+ (FCVT_D_WU, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2780_
+ _s1342_) >>= fun w__373 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__373 with
+ | Some
+ (FCVT_D_WU, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_WU))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2797_
+ _s1342_) >>= fun w__376 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__376 with
+ | Some
+ (FCVT_L_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2797_
+ _s1342_) >>= fun w__377 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__377 with
+ | Some
+ (FCVT_L_D, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_L_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2814_
+ _s1342_) >>= fun w__380 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__380 with
+ | Some
+ (FCVT_LU_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2814_
+ _s1342_) >>= fun w__381 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__381 with
+ | Some
+ (FCVT_LU_D, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_LU_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2831_
+ _s1342_) >>= fun w__384 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__384 with
+ | Some
+ (FCVT_D_L, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2831_
+ _s1342_) >>= fun w__385 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__385 with
+ | Some
+ (FCVT_D_L, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_L))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2848_
+ _s1342_) >>= fun w__388 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__388 with
+ | Some
+ (FCVT_D_LU, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2848_
+ _s1342_) >>= fun w__389 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__389 with
+ | Some
+ (FCVT_D_LU, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_LU))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2865_
+ _s1342_) >>= fun w__392 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__392 with
+ | Some
+ (FCVT_S_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2865_
+ _s1342_) >>= fun w__393 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__393 with
+ | Some
+ (FCVT_S_D, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_S_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2882_
+ _s1342_) >>= fun w__396 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__396 with
+ | Some
+ (FCVT_D_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2882_
+ _s1342_) >>= fun w__397 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__397 with
+ | Some
+ (FCVT_D_S, rd, rs1, rm) =>
+ returnm (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_S))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2899_
+ _s1342_) >>= fun w__400 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__400 with
+ | Some
+ (FSGNJ_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2899_
+ _s1342_) >>= fun w__401 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__401 with
+ | Some
+ (FSGNJ_D, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJ_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2916_
+ _s1342_) >>= fun w__404 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__404 with
+ | Some
+ (FSGNJN_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2916_
+ _s1342_) >>= fun w__405 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__405 with
+ | Some
+ (FSGNJN_D, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJN_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2933_
+ _s1342_) >>= fun w__408 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__408 with
+ | Some
+ (FSGNJX_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2933_
+ _s1342_) >>= fun w__409 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__409 with
+ | Some
+ (FSGNJX_D, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJX_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2950_
+ _s1342_) >>= fun w__412 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__412 with
+ | Some
+ (FMIN_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2950_
+ _s1342_) >>= fun w__413 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__413 with
+ | Some
+ (FMIN_D, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FMIN_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2967_
+ _s1342_) >>= fun w__416 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__416 with
+ | Some
+ (FMAX_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2967_
+ _s1342_) >>= fun w__417 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__417 with
+ | Some
+ (FMAX_D, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FMAX_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s2984_
+ _s1342_) >>= fun w__420 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__420 with
+ | Some
+ (FEQ_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s2984_
+ _s1342_) >>= fun w__421 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__421 with
+ | Some
+ (FEQ_D, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FEQ_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3001_
+ _s1342_) >>= fun w__424 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__424 with
+ | Some
+ (FLT_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s3001_
+ _s1342_) >>= fun w__425 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__425 with
+ | Some
+ (FLT_D, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FLT_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3018_
+ _s1342_) >>= fun w__428 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__428 with
+ | Some
+ (FLE_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s3018_
+ _s1342_) >>= fun w__429 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__429 with
+ | Some
+ (FLE_D, rd, rs1, rs2) =>
+ returnm (F_BIN_TYPE_D
+ (rs2, rs1, rd, FLE_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3035_
+ _s1342_) >>= fun w__432 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (if
+ match w__432 with
+ | Some
+ (FMV_X_D, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s3035_
+ _s1342_) >>= fun w__433 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (match w__433 with
+ | Some
+ (FMV_X_D, rd, rs1) =>
+ returnm (F_UN_TYPE_D
+ (rs1, rd, FMV_X_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3048_
+ _s1342_) >>= fun w__436 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (if
+ match w__436 with
+ | Some
+ (FMV_D_X, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s3048_
+ _s1342_) >>= fun w__437 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (match w__437 with
+ | Some
+ (FMV_D_X, rd, rs1) =>
+ returnm (F_UN_TYPE_D
+ (rs1, rd, FMV_D_X))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3061_
+ _s1342_) >>= fun w__440 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (if
+ match w__440 with
+ | Some
+ (FCLASS_D, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s3061_
+ _s1342_) >>= fun w__441 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (match w__441 with
+ | Some
+ (FCLASS_D, rd, rs1) =>
+ returnm (F_UN_TYPE_D
+ (rs1, rd, FCLASS_D))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3074_
+ _s1342_) >>= fun w__444 : option ((mword 5 * mword 6)) =>
+ (if
+ match w__444 with
+ | Some
+ (rd, uimm) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s3074_
+ _s1342_) >>= fun w__445 : option ((mword 5 * mword 6)) =>
+ (match w__445 with
+ | Some
+ (rd, uimm) =>
+ returnm (C_FLDSP
+ (uimm, rd))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3086_
+ _s1342_) >>= fun w__448 : option ((mword 5 * mword 6)) =>
+ (if
+ match w__448 with
+ | Some
+ (rs2, uimm) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s3086_
+ _s1342_) >>= fun w__449 : option ((mword 5 * mword 6)) =>
+ (match w__449 with
+ | Some
+ (rs2, uimm) =>
+ returnm (C_FSDSP
+ (uimm, rs2))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3098_
+ _s1342_) >>= fun w__452 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if
+ match w__452 with
+ | Some
+ (rdc, rsc, uimm) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s3098_
+ _s1342_) >>= fun w__453 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__453 with
+ | Some
+ (rdc, rsc, uimm) =>
+ returnm (C_FLD
+ (uimm, rsc, rdc))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3114_
+ _s1342_) >>= fun w__456 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if
+ match w__456 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s3114_
+ _s1342_) >>= fun w__457 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__457 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ returnm (C_FSD
+ (uimm, rsc1, rsc2))
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3130_
+ _s1342_) >>= fun w__460 : option (mword 32) =>
+ (if
+ match w__460 with
+ | Some
+ s =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s3130_
+ _s1342_) >>= fun w__461 : option (mword 32) =>
+ (match w__461 with
+ | Some
+ s =>
+ returnm (ILLEGAL
+ s)
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ (_s3138_
+ _s1342_) >>= fun w__464 : option (mword 16) =>
+ (if
+ match w__464 with
+ | Some
+ s =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s3138_
+ _s1342_) >>= fun w__465 : option (mword 16) =>
+ (match w__465 with
+ | Some
+ s =>
+ returnm (C_ILLEGAL
+ s)
+ | _ =>
+ exit tt
+ : M (ast)
+ end)
+ : M (ast)
+ else
+ assert_exp' false "Pattern match failure at unknown location" >>= fun _ =>
+ exit tt)
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
+ : M (ast))
: M (ast))
: M (ast))
: M (ast))
@@ -29761,6053 +42984,11357 @@ Definition assembly_backwards (arg_ : string)
: M (ast))
: M (ast).
-Definition assembly_forwards_matches (arg_ : ast)
-: bool :=
-
+Definition assembly_forwards_matches (arg_ : ast) : bool :=
match arg_ with
- | UTYPE ((imm, rd, op)) => true
- | RISCV_JAL ((imm, rd)) => true
- | RISCV_JALR ((imm, rs1, rd)) => true
- | BTYPE ((imm, rs2, rs1, op)) => true
- | ITYPE ((imm, rs1, rd, op)) => true
- | SHIFTIOP ((shamt, rs1, rd, op)) => true
- | RTYPE ((rs2, rs1, rd, op)) => true
- | LOAD ((imm, rs1, rd, is_unsigned, size, aq, rl)) => true
- | STORE ((imm, rs2, rs1, size, aq, rl)) => true
- | ADDIW ((imm, rs1, rd)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | SHIFTW ((shamt, rs1, rd, op)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | RTYPEW ((rs2, rs1, rd, op)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | SHIFTIWOP ((shamt, rs1, rd, op)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | FENCE ((pred, succ)) => true
- | FENCE_TSO ((pred, succ)) => true
- | FENCEI (tt) => true
- | ECALL (tt) => true
- | MRET (tt) => true
- | SRET (tt) => true
- | EBREAK (tt) => true
- | WFI (tt) => true
- | SFENCE_VMA ((rs1, rs2)) => true
- | LOADRES ((aq, rl, rs1, size, rd)) => true
- | STORECON ((aq, rl, rs2, rs1, size, rd)) => true
- | AMO ((op, aq, rl, rs2, rs1, width, rd)) => true
- | C_NOP (tt) => true
- | C_ADDI4SPN ((rdc, nzimm)) =>
- if ((neq_vec nzimm (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0] : mword 8))) then true else false
- | C_LW ((uimm, rsc, rdc)) => true
- | C_LD ((uimm, rsc, rdc)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_SW ((uimm, rsc1, rsc2)) => true
- | C_SD ((uimm, rsc1, rsc2)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_ADDI ((nzi, rsd)) =>
- if ((andb (neq_vec nzi (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- : bool))) then
- true
- else false
- | C_JAL (imm) => if sumbool_of_bool ((Z.eqb 64 32)) then true else false
- | C_ADDIW ((imm, rsd)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_LI ((imm, rd)) =>
- if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg)))))
- then
- true
- else false
- | C_ADDI16SP (imm) =>
- if ((neq_vec imm (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then true else false
- | C_LUI ((imm, rd)) =>
- if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg))))
- ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd))
- (projT1
- (regidx_to_regno sp))))
- (neq_vec imm (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6)))
- : bool))) then
- true
- else false
- | C_SRLI ((shamt, rsd)) =>
- if ((neq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then true else false
- | C_SRAI ((shamt, rsd)) =>
- if ((neq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))) then true else false
- | C_ANDI ((imm, rsd)) => true
- | C_SUB ((rsd, rs2)) => true
- | C_XOR ((rsd, rs2)) => true
- | C_OR ((rsd, rs2)) => true
- | C_AND ((rsd, rs2)) => true
- | C_SUBW ((rsd, rs2)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_ADDW ((rsd, rs2)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_J (imm) => true
- | C_BEQZ ((imm, rs)) => true
- | C_BNEZ ((imm, rs)) => true
- | C_SLLI ((shamt, rsd)) =>
- if ((andb (neq_vec shamt (vec_of_bits [B0;B0;B0;B0;B0;B0] : mword 6))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rsd)) (projT1 (regidx_to_regno zreg))))
- : bool))) then
- true
- else false
- | C_LWSP ((uimm, rd)) =>
- if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg)))))
- then
- true
- else false
- | C_LDSP ((uimm, rd)) =>
- if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg))))
- (Z.eqb 64 64))) then
- true
- else false
- | C_SWSP ((uimm, rd)) => true
- | C_SDSP ((uimm, rs2)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | C_JR (rs1) =>
- if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rs1)) (projT1 (regidx_to_regno zreg)))))
- then
- true
- else false
- | C_JALR (rs1) =>
- if sumbool_of_bool ((projT1
- (neq_int (projT1 (regidx_to_regno rs1)) (projT1 (regidx_to_regno zreg)))))
- then
- true
- else false
- | C_MV ((rd, rs2)) =>
- if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rd)) (projT1 (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))) then
+ | UTYPE (imm, rd, op) => true
+ | RISCV_JAL (imm, rd) => true
+ | RISCV_JALR (imm, rs1, rd) => true
+ | BTYPE (imm, rs2, rs1, op) => true
+ | ITYPE (imm, rs1, rd, op) => true
+ | SHIFTIOP (shamt, rs1, rd, op) => true
+ | RTYPE (rs2, rs1, rd, op) => true
+ | LOAD (imm, rs1, rd, is_unsigned, size, aq, rl) => true
+ | STORE (imm, rs2, rs1, size, aq, rl) => true
+ | ADDIW (imm, rs1, rd) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | SHIFTW (shamt, rs1, rd, op) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | RTYPEW (rs2, rs1, rd, op) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | SHIFTIWOP (shamt, rs1, rd, op) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | FENCE (pred, succ) => true
+ | FENCE_TSO (pred, succ) => true
+ | FENCEI tt => true
+ | ECALL tt => true
+ | MRET tt => true
+ | SRET tt => true
+ | EBREAK tt => true
+ | WFI tt => true
+ | SFENCE_VMA (rs1, rs2) => true
+ | LOADRES (aq, rl, rs1, size, rd) => true
+ | STORECON (aq, rl, rs2, rs1, size, rd) => true
+ | AMO (op, aq, rl, rs2, rs1, width, rd) => true
+ | C_NOP tt => true
+ | C_ADDI4SPN (rdc, nzimm) => if neq_vec nzimm (Ox"00" : mword 8) then true else false
+ | C_LW (uimm, rsc, rdc) => true
+ | C_LD (uimm, rsc, rdc) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | C_SW (uimm, rsc1, rsc2) => true
+ | C_SD (uimm, rsc1, rsc2) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | C_ADDI (nzi, rsd) =>
+ if andb (neq_vec nzi ('b"000000" : mword 6)) (neq_vec rsd zreg) then true else false
+ | C_JAL imm => if sumbool_of_bool (Z.eqb 64 32) then true else false
+ | C_ADDIW (imm, rsd) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | C_LI (imm, rd) => if neq_vec rd zreg then true else false
+ | C_ADDI16SP imm => if neq_vec imm ('b"000000" : mword 6) then true else false
+ | C_LUI (imm, rd) =>
+ if andb (neq_vec rd zreg) (andb (neq_vec rd sp) (neq_vec imm ('b"000000" : mword 6))) then
true
else false
- | C_EBREAK (tt) => true
- | C_ADD ((rsd, rs2)) =>
- if sumbool_of_bool ((andb
- (projT1
- (neq_int (projT1 (regidx_to_regno rsd))
- (projT1
- (regidx_to_regno zreg))))
- ((projT1
- (neq_int (projT1 (regidx_to_regno rs2))
- (projT1
- (regidx_to_regno zreg))))
- : bool))) then
- true
- else false
- | MUL ((rs2, rs1, rd, high, signed1, signed2)) => true
- | DIV ((rs2, rs1, rd, s)) => true
- | REM ((rs2, rs1, rd, s)) => true
- | MULW ((rs2, rs1, rd)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | DIVW ((rs2, rs1, rd, s)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | REMW ((rs2, rs1, rd, s)) => if sumbool_of_bool ((Z.eqb 64 64)) then true else false
- | CSR ((csr, rs1, rd, true, op)) => true
- | CSR ((csr, rs1, rd, false, op)) => true
- | URET (tt) => true
- | ILLEGAL (s) => true
- | C_ILLEGAL (s) => true
+ | C_SRLI (shamt, rsd) => if neq_vec shamt ('b"000000" : mword 6) then true else false
+ | C_SRAI (shamt, rsd) => if neq_vec shamt ('b"000000" : mword 6) then true else false
+ | C_ANDI (imm, rsd) => true
+ | C_SUB (rsd, rs2) => true
+ | C_XOR (rsd, rs2) => true
+ | C_OR (rsd, rs2) => true
+ | C_AND (rsd, rs2) => true
+ | C_SUBW (rsd, rs2) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | C_ADDW (rsd, rs2) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | C_J imm => true
+ | C_BEQZ (imm, rs) => true
+ | C_BNEZ (imm, rs) => true
+ | C_SLLI (shamt, rsd) =>
+ if andb (neq_vec shamt ('b"000000" : mword 6)) (neq_vec rsd zreg) then true else false
+ | C_LWSP (uimm, rd) => if neq_vec rd zreg then true else false
+ | C_LDSP (uimm, rd) =>
+ if sumbool_of_bool (andb (neq_vec rd zreg) (Z.eqb 64 64)) then true else false
+ | C_SWSP (uimm, rd) => true
+ | C_SDSP (uimm, rs2) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | C_JR rs1 => if neq_vec rs1 zreg then true else false
+ | C_JALR rs1 => if neq_vec rs1 zreg then true else false
+ | C_MV (rd, rs2) => if andb (neq_vec rd zreg) (neq_vec rs2 zreg) then true else false
+ | C_EBREAK tt => true
+ | C_ADD (rsd, rs2) => if andb (neq_vec rsd zreg) (neq_vec rs2 zreg) then true else false
+ | MUL (rs2, rs1, rd, high, signed1, signed2) => true
+ | DIV (rs2, rs1, rd, s) => true
+ | REM (rs2, rs1, rd, s) => true
+ | MULW (rs2, rs1, rd) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | DIVW (rs2, rs1, rd, s) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | REMW (rs2, rs1, rd, s) => if sumbool_of_bool (Z.eqb 64 64) then true else false
+ | CSR (csr, rs1, rd, true, op) => true
+ | CSR (csr, rs1, rd, false, op) => true
+ | URET tt => true
+ | LOAD_FP (imm, rs1, rd, width) => true
+ | STORE_FP (imm, rs2, rs1, width) => true
+ | F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, op) => true
+ | F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, op) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FSQRT_S) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_W_S) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_WU_S) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_W) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_WU) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_L_S) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_LU_S) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_L) => true
+ | F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_LU) => true
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJ_S) => true
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJN_S) => true
+ | F_BIN_TYPE_S (rs2, rs1, rd, FSGNJX_S) => true
+ | F_BIN_TYPE_S (rs2, rs1, rd, FMIN_S) => true
+ | F_BIN_TYPE_S (rs2, rs1, rd, FMAX_S) => true
+ | F_BIN_TYPE_S (rs2, rs1, rd, FEQ_S) => true
+ | F_BIN_TYPE_S (rs2, rs1, rd, FLT_S) => true
+ | F_BIN_TYPE_S (rs2, rs1, rd, FLE_S) => true
+ | F_UN_TYPE_S (rs1, rd, FMV_X_W) => true
+ | F_UN_TYPE_S (rs1, rd, FMV_W_X) => true
+ | F_UN_TYPE_S (rs1, rd, FCLASS_S) => true
+ | C_FLWSP (imm, rd) => if sumbool_of_bool (Z.eqb 64 32) then true else false
+ | C_FSWSP (uimm, rd) => if sumbool_of_bool (Z.eqb 64 32) then true else false
+ | C_FLW (uimm, rsc, rdc) => if sumbool_of_bool (Z.eqb 64 32) then true else false
+ | C_FSW (uimm, rsc1, rsc2) => if sumbool_of_bool (Z.eqb 64 32) then true else false
+ | F_MADD_TYPE_D (rs3, rs2, rs1, rm, rd, op) => true
+ | F_BIN_RM_TYPE_D (rs2, rs1, rm, rd, op) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FSQRT_D) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_W_D) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_WU_D) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_W) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_WU) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_L_D) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_LU_D) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_L) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_LU) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_S_D) => true
+ | F_UN_RM_TYPE_D (rs1, rm, rd, FCVT_D_S) => true
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJ_D) => true
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJN_D) => true
+ | F_BIN_TYPE_D (rs2, rs1, rd, FSGNJX_D) => true
+ | F_BIN_TYPE_D (rs2, rs1, rd, FMIN_D) => true
+ | F_BIN_TYPE_D (rs2, rs1, rd, FMAX_D) => true
+ | F_BIN_TYPE_D (rs2, rs1, rd, FEQ_D) => true
+ | F_BIN_TYPE_D (rs2, rs1, rd, FLT_D) => true
+ | F_BIN_TYPE_D (rs2, rs1, rd, FLE_D) => true
+ | F_UN_TYPE_D (rs1, rd, FMV_X_D) => true
+ | F_UN_TYPE_D (rs1, rd, FMV_D_X) => true
+ | F_UN_TYPE_D (rs1, rd, FCLASS_D) => true
+ | C_FLDSP (uimm, rd) => if sumbool_of_bool (orb (Z.eqb 64 32) (Z.eqb 64 64)) then true else false
+ | C_FSDSP (uimm, rs2) =>
+ if sumbool_of_bool (orb (Z.eqb 64 32) (Z.eqb 64 64)) then true else false
+ | C_FLD (uimm, rsc, rdc) =>
+ if sumbool_of_bool (orb (Z.eqb 64 32) (Z.eqb 64 64)) then true else false
+ | C_FSD (uimm, rsc1, rsc2) =>
+ if sumbool_of_bool (orb (Z.eqb 64 32) (Z.eqb 64 64)) then true else false
+ | ILLEGAL s => true
+ | C_ILLEGAL s => true
end.
-Definition _s2549_ (_s2550_ : string)
-: M (option (mword 16)) :=
-
- let _s2551_ := _s2550_ in
- (if ((string_startswith _s2551_ "c.illegal")) then
- (match (string_drop _s2551_ (projT1 (string_length "c.illegal"))) with
- | _s2552_ =>
- (spc_matches_prefix _s2552_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s2553_ _)) =>
- match (string_drop _s2552_ _s2553_) with
- | _s2554_ =>
- match (hex_bits_16_matches_prefix _s2554_) with
- | Some ((s, existT _ _s2555_ _)) =>
- let p0_ := string_drop _s2554_ _s2555_ in
- if ((generic_eq p0_ "")) then Some (s)
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 16))
+Definition _s4954_ (_s4955_ : string) : M (option (mword 16)) :=
+ let _s4956_ := _s4955_ in
+ (if string_startswith _s4956_ "c.illegal" then
+ (match (string_drop _s4956_ (projT1 (string_length "c.illegal"))) with
+ | _s4957_ =>
+ (spc_matches_prefix _s4957_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s4958_ _) =>
+ match (string_drop _s4957_ _s4958_) with
+ | _s4959_ =>
+ match (hex_bits_16_matches_prefix _s4959_) with
+ | Some (s, existT _ _s4960_ _) =>
+ let p0_ := string_drop _s4959_ _s4960_ in
+ if generic_eq p0_ "" then Some s
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 16))
- else returnm (None : option (mword 16)))
+ else returnm None)
: M (option (mword 16)).
-Definition _s2541_ (_s2542_ : string)
-: M (option (mword 32)) :=
-
- let _s2543_ := _s2542_ in
- (if ((string_startswith _s2543_ "illegal")) then
- (match (string_drop _s2543_ (projT1 (string_length "illegal"))) with
- | _s2544_ =>
- (spc_matches_prefix _s2544_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s2545_ _)) =>
- match (string_drop _s2544_ _s2545_) with
- | _s2546_ =>
- match (hex_bits_32_matches_prefix _s2546_) with
- | Some ((s, existT _ _s2547_ _)) =>
- let p0_ := string_drop _s2546_ _s2547_ in
- if ((generic_eq p0_ "")) then Some (s)
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 32))
+Definition _s4946_ (_s4947_ : string) : M (option (mword 32)) :=
+ let _s4948_ := _s4947_ in
+ (if string_startswith _s4948_ "illegal" then
+ (match (string_drop _s4948_ (projT1 (string_length "illegal"))) with
+ | _s4949_ =>
+ (spc_matches_prefix _s4949_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s4950_ _) =>
+ match (string_drop _s4949_ _s4950_) with
+ | _s4951_ =>
+ match (hex_bits_32_matches_prefix _s4951_) with
+ | Some (s, existT _ _s4952_ _) =>
+ let p0_ := string_drop _s4951_ _s4952_ in
+ if generic_eq p0_ "" then Some s
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 32))
- else returnm (None : option (mword 32)))
+ else returnm None)
: M (option (mword 32)).
-Definition _s2524_ (_s2525_ : string)
-: M (option ((csrop * mword 5 * mword 5 * mword 12))) :=
-
- (match _s2525_ with
- | _s2526_ =>
- (csr_mnemonic_matches_prefix _s2526_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=
+Definition _s4930_ (_s4931_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s4932_ := _s4931_ in
+ (if string_startswith _s4932_ "c.fsd" then
+ (match (string_drop _s4932_ (projT1 (string_length "c.fsd"))) with
+ | _s4933_ =>
+ (spc_matches_prefix _s4933_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s4934_ _) =>
+ (match (string_drop _s4933_ _s4934_) with
+ | _s4935_ =>
+ (creg_name_matches_prefix _s4935_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rsc1, existT _ _s4936_ _) =>
+ (match (string_drop _s4935_ _s4936_) with
+ | _s4937_ =>
+ (sep_matches_prefix _s4937_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s4938_ _) =>
+ (match (string_drop _s4937_ _s4938_) with
+ | _s4939_ =>
+ (creg_name_matches_prefix _s4939_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc2, existT _ _s4940_ _) =>
+ (match (string_drop _s4939_ _s4940_) with
+ | _s4941_ =>
+ (sep_matches_prefix _s4941_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s4942_ _) =>
+ match (string_drop _s4941_ _s4942_) with
+ | _s4943_ =>
+ match (hex_bits_8_matches_prefix _s4943_) with
+ | Some (v__1358, existT _ _s4944_ _) =>
+ if eq_vec (subrange_vec_dec v__1358 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1358 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1358 7 3 in
+ let p0_ := string_drop _s4943_ _s4944_ in
+ if generic_eq p0_ "" then
+ Some (rsc1, rsc2, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5))).
+
+Definition _s4914_ (_s4915_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s4916_ := _s4915_ in
+ (if string_startswith _s4916_ "c.fld" then
+ (match (string_drop _s4916_ (projT1 (string_length "c.fld"))) with
+ | _s4917_ =>
+ (spc_matches_prefix _s4917_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s4918_ _) =>
+ (match (string_drop _s4917_ _s4918_) with
+ | _s4919_ =>
+ (creg_name_matches_prefix _s4919_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rdc, existT _ _s4920_ _) =>
+ (match (string_drop _s4919_ _s4920_) with
+ | _s4921_ =>
+ (sep_matches_prefix _s4921_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s4922_ _) =>
+ (match (string_drop _s4921_ _s4922_) with
+ | _s4923_ =>
+ (creg_name_matches_prefix _s4923_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc, existT _ _s4924_ _) =>
+ (match (string_drop _s4923_ _s4924_) with
+ | _s4925_ =>
+ (sep_matches_prefix _s4925_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s4926_ _) =>
+ match (string_drop _s4925_ _s4926_) with
+ | _s4927_ =>
+ match (hex_bits_8_matches_prefix _s4927_) with
+ | Some (v__1360, existT _ _s4928_ _) =>
+ if eq_vec (subrange_vec_dec v__1360 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1360 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1360 7 3 in
+ let p0_ := string_drop _s4927_ _s4928_ in
+ if generic_eq p0_ "" then
+ Some (rdc, rsc, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5))).
+
+Definition _s4902_ (_s4903_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s4904_ := _s4903_ in
+ (if string_startswith _s4904_ "c.fsdsp" then
+ (match (string_drop _s4904_ (projT1 (string_length "c.fsdsp"))) with
+ | _s4905_ =>
+ (spc_matches_prefix _s4905_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s4906_ _) =>
+ (match (string_drop _s4905_ _s4906_) with
+ | _s4907_ =>
+ (reg_name_matches_prefix _s4907_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rs2, existT _ _s4908_ _) =>
+ (match (string_drop _s4907_ _s4908_) with
+ | _s4909_ =>
+ (sep_matches_prefix _s4909_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s4910_ _) =>
+ match (string_drop _s4909_ _s4910_) with
+ | _s4911_ =>
+ match (hex_bits_6_matches_prefix _s4911_) with
+ | Some (uimm, existT _ _s4912_ _) =>
+ let p0_ := string_drop _s4911_ _s4912_ in
+ if generic_eq p0_ "" then Some (rs2, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6))).
+
+Definition _s4890_ (_s4891_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s4892_ := _s4891_ in
+ (if string_startswith _s4892_ "c.fldsp" then
+ (match (string_drop _s4892_ (projT1 (string_length "c.fldsp"))) with
+ | _s4893_ =>
+ (spc_matches_prefix _s4893_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s4894_ _) =>
+ (match (string_drop _s4893_ _s4894_) with
+ | _s4895_ =>
+ (reg_name_matches_prefix _s4895_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s4896_ _) =>
+ (match (string_drop _s4895_ _s4896_) with
+ | _s4897_ =>
+ (sep_matches_prefix _s4897_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s4898_ _) =>
+ match (string_drop _s4897_ _s4898_) with
+ | _s4899_ =>
+ match (hex_bits_6_matches_prefix _s4899_) with
+ | Some (uimm, existT _ _s4900_ _) =>
+ let p0_ := string_drop _s4899_ _s4900_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6))).
+
+Definition _s4877_ (_s4878_ : string) : M (option ((f_un_op_D * mword 5 * mword 5))) :=
+ (match _s4878_ with
+ | _s4879_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s4879_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2527_ _)) =>
- (match (string_drop _s2526_ _s2527_) with
- | _s2528_ =>
- (spc_matches_prefix _s2528_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (FCLASS_D, existT _ _s4880_ _) =>
+ (match (string_drop _s4879_ _s4880_) with
+ | _s4881_ =>
+ (spc_matches_prefix _s4881_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2529_ _)) =>
- (match (string_drop _s2528_ _s2529_) with
- | _s2530_ =>
- (reg_name_matches_prefix _s2530_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s4882_ _) =>
+ (match (string_drop _s4881_ _s4882_) with
+ | _s4883_ =>
+ (reg_name_matches_prefix _s4883_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2531_ _)) =>
- (match (string_drop _s2530_ _s2531_) with
- | _s2532_ =>
- (sep_matches_prefix _s2532_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s4884_ _) =>
+ (match (string_drop _s4883_ _s4884_) with
+ | _s4885_ =>
+ (sep_matches_prefix _s4885_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2533_ _)) =>
- (match (string_drop _s2532_ _s2533_) with
- | _s2534_ =>
- (reg_name_matches_prefix _s2534_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s4886_ _) =>
+ (match (string_drop _s4885_ _s4886_) with
+ | _s4887_ =>
+ (freg_name_matches_prefix _s4887_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s4888_ _) =>
+ let p0_ := string_drop _s4887_ _s4888_ in
+ if generic_eq p0_ "" then
+ Some (FCLASS_D, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5))).
+
+Definition _s4864_ (_s4865_ : string) : M (option ((f_un_op_D * mword 5 * mword 5))) :=
+ (match _s4865_ with
+ | _s4866_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s4866_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_D_X, existT _ _s4867_ _) =>
+ (match (string_drop _s4866_ _s4867_) with
+ | _s4868_ =>
+ (spc_matches_prefix _s4868_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4869_ _) =>
+ (match (string_drop _s4868_ _s4869_) with
+ | _s4870_ =>
+ (freg_name_matches_prefix _s4870_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4871_ _) =>
+ (match (string_drop _s4870_ _s4871_) with
+ | _s4872_ =>
+ (sep_matches_prefix _s4872_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4873_ _) =>
+ (match (string_drop _s4872_ _s4873_) with
+ | _s4874_ =>
+ (reg_name_matches_prefix _s4874_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s4875_ _) =>
+ let p0_ := string_drop _s4874_ _s4875_ in
+ if generic_eq p0_ "" then Some (FMV_D_X, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5))).
+
+Definition _s4851_ (_s4852_ : string) : M (option ((f_un_op_D * mword 5 * mword 5))) :=
+ (match _s4852_ with
+ | _s4853_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s4853_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_X_D, existT _ _s4854_ _) =>
+ (match (string_drop _s4853_ _s4854_) with
+ | _s4855_ =>
+ (spc_matches_prefix _s4855_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4856_ _) =>
+ (match (string_drop _s4855_ _s4856_) with
+ | _s4857_ =>
+ (reg_name_matches_prefix _s4857_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4858_ _) =>
+ (match (string_drop _s4857_ _s4858_) with
+ | _s4859_ =>
+ (sep_matches_prefix _s4859_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4860_ _) =>
+ (match (string_drop _s4859_ _s4860_) with
+ | _s4861_ =>
+ (freg_name_matches_prefix _s4861_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s4862_ _) =>
+ let p0_ := string_drop _s4861_ _s4862_ in
+ if generic_eq p0_ "" then Some (FMV_X_D, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5))).
+
+Definition _s4834_ (_s4835_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s4835_ with
+ | _s4836_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s4836_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLE_D, existT _ _s4837_ _) =>
+ (match (string_drop _s4836_ _s4837_) with
+ | _s4838_ =>
+ (spc_matches_prefix _s4838_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4839_ _) =>
+ (match (string_drop _s4838_ _s4839_) with
+ | _s4840_ =>
+ (reg_name_matches_prefix _s4840_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4841_ _) =>
+ (match (string_drop _s4840_ _s4841_) with
+ | _s4842_ =>
+ (sep_matches_prefix _s4842_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4843_ _) =>
+ (match (string_drop _s4842_ _s4843_) with
+ | _s4844_ =>
+ (freg_name_matches_prefix _s4844_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2535_ _)) =>
- (match (string_drop _s2534_ _s2535_) with
- | _s2536_ =>
- (sep_matches_prefix _s2536_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s4845_ _) =>
+ (match (string_drop _s4844_ _s4845_) with
+ | _s4846_ =>
+ (sep_matches_prefix _s4846_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2537_ _)) =>
- (match (string_drop _s2536_ _s2537_) with
- | _s2538_ =>
- (csr_name_map_matches_prefix _s2538_) >>= fun w__6 : option ((mword 12 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s4847_ _) =>
+ (match (string_drop _s4846_ _s4847_) with
+ | _s4848_ =>
+ (freg_name_matches_prefix _s4848_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((csr, existT _ _s2539_ _)) =>
- let p0_ :=
- string_drop _s2538_ _s2539_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, csr))
- else None
- | _ => None
- end)
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4849_ _) =>
+ let p0_ :=
+ string_drop _s4848_ _s4849_ in
+ if generic_eq p0_ "" then
+ Some (FLE_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s4817_ (_s4818_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s4818_ with
+ | _s4819_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s4819_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLT_D, existT _ _s4820_ _) =>
+ (match (string_drop _s4819_ _s4820_) with
+ | _s4821_ =>
+ (spc_matches_prefix _s4821_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4822_ _) =>
+ (match (string_drop _s4821_ _s4822_) with
+ | _s4823_ =>
+ (reg_name_matches_prefix _s4823_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4824_ _) =>
+ (match (string_drop _s4823_ _s4824_) with
+ | _s4825_ =>
+ (sep_matches_prefix _s4825_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4826_ _) =>
+ (match (string_drop _s4825_ _s4826_) with
+ | _s4827_ =>
+ (freg_name_matches_prefix _s4827_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4828_ _) =>
+ (match (string_drop _s4827_ _s4828_) with
+ | _s4829_ =>
+ (sep_matches_prefix _s4829_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4830_ _) =>
+ (match (string_drop _s4829_ _s4830_) with
+ | _s4831_ =>
+ (freg_name_matches_prefix _s4831_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4832_ _) =>
+ let p0_ :=
+ string_drop _s4831_ _s4832_ in
+ if generic_eq p0_ "" then
+ Some (FLT_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s4800_ (_s4801_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s4801_ with
+ | _s4802_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s4802_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FEQ_D, existT _ _s4803_ _) =>
+ (match (string_drop _s4802_ _s4803_) with
+ | _s4804_ =>
+ (spc_matches_prefix _s4804_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4805_ _) =>
+ (match (string_drop _s4804_ _s4805_) with
+ | _s4806_ =>
+ (reg_name_matches_prefix _s4806_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4807_ _) =>
+ (match (string_drop _s4806_ _s4807_) with
+ | _s4808_ =>
+ (sep_matches_prefix _s4808_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4809_ _) =>
+ (match (string_drop _s4808_ _s4809_) with
+ | _s4810_ =>
+ (freg_name_matches_prefix _s4810_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4811_ _) =>
+ (match (string_drop _s4810_ _s4811_) with
+ | _s4812_ =>
+ (sep_matches_prefix _s4812_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4813_ _) =>
+ (match (string_drop _s4812_ _s4813_) with
+ | _s4814_ =>
+ (freg_name_matches_prefix _s4814_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4815_ _) =>
+ let p0_ :=
+ string_drop _s4814_ _s4815_ in
+ if generic_eq p0_ "" then
+ Some (FEQ_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s4783_ (_s4784_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s4784_ with
+ | _s4785_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s4785_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMAX_D, existT _ _s4786_ _) =>
+ (match (string_drop _s4785_ _s4786_) with
+ | _s4787_ =>
+ (spc_matches_prefix _s4787_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4788_ _) =>
+ (match (string_drop _s4787_ _s4788_) with
+ | _s4789_ =>
+ (freg_name_matches_prefix _s4789_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4790_ _) =>
+ (match (string_drop _s4789_ _s4790_) with
+ | _s4791_ =>
+ (sep_matches_prefix _s4791_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4792_ _) =>
+ (match (string_drop _s4791_ _s4792_) with
+ | _s4793_ =>
+ (freg_name_matches_prefix _s4793_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4794_ _) =>
+ (match (string_drop _s4793_ _s4794_) with
+ | _s4795_ =>
+ (sep_matches_prefix _s4795_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4796_ _) =>
+ (match (string_drop _s4795_ _s4796_) with
+ | _s4797_ =>
+ (freg_name_matches_prefix _s4797_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4798_ _) =>
+ let p0_ :=
+ string_drop _s4797_ _s4798_ in
+ if generic_eq p0_ "" then
+ Some (FMAX_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s4766_ (_s4767_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s4767_ with
+ | _s4768_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s4768_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMIN_D, existT _ _s4769_ _) =>
+ (match (string_drop _s4768_ _s4769_) with
+ | _s4770_ =>
+ (spc_matches_prefix _s4770_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4771_ _) =>
+ (match (string_drop _s4770_ _s4771_) with
+ | _s4772_ =>
+ (freg_name_matches_prefix _s4772_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4773_ _) =>
+ (match (string_drop _s4772_ _s4773_) with
+ | _s4774_ =>
+ (sep_matches_prefix _s4774_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4775_ _) =>
+ (match (string_drop _s4774_ _s4775_) with
+ | _s4776_ =>
+ (freg_name_matches_prefix _s4776_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4777_ _) =>
+ (match (string_drop _s4776_ _s4777_) with
+ | _s4778_ =>
+ (sep_matches_prefix _s4778_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4779_ _) =>
+ (match (string_drop _s4778_ _s4779_) with
+ | _s4780_ =>
+ (freg_name_matches_prefix _s4780_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4781_ _) =>
+ let p0_ :=
+ string_drop _s4780_ _s4781_ in
+ if generic_eq p0_ "" then
+ Some (FMIN_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s4749_ (_s4750_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s4750_ with
+ | _s4751_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s4751_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJX_D, existT _ _s4752_ _) =>
+ (match (string_drop _s4751_ _s4752_) with
+ | _s4753_ =>
+ (spc_matches_prefix _s4753_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4754_ _) =>
+ (match (string_drop _s4753_ _s4754_) with
+ | _s4755_ =>
+ (freg_name_matches_prefix _s4755_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4756_ _) =>
+ (match (string_drop _s4755_ _s4756_) with
+ | _s4757_ =>
+ (sep_matches_prefix _s4757_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4758_ _) =>
+ (match (string_drop _s4757_ _s4758_) with
+ | _s4759_ =>
+ (freg_name_matches_prefix _s4759_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4760_ _) =>
+ (match (string_drop _s4759_ _s4760_) with
+ | _s4761_ =>
+ (sep_matches_prefix _s4761_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4762_ _) =>
+ (match (string_drop _s4761_ _s4762_) with
+ | _s4763_ =>
+ (freg_name_matches_prefix _s4763_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4764_ _) =>
+ let p0_ :=
+ string_drop _s4763_ _s4764_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJX_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s4732_ (_s4733_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s4733_ with
+ | _s4734_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s4734_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJN_D, existT _ _s4735_ _) =>
+ (match (string_drop _s4734_ _s4735_) with
+ | _s4736_ =>
+ (spc_matches_prefix _s4736_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4737_ _) =>
+ (match (string_drop _s4736_ _s4737_) with
+ | _s4738_ =>
+ (freg_name_matches_prefix _s4738_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4739_ _) =>
+ (match (string_drop _s4738_ _s4739_) with
+ | _s4740_ =>
+ (sep_matches_prefix _s4740_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4741_ _) =>
+ (match (string_drop _s4740_ _s4741_) with
+ | _s4742_ =>
+ (freg_name_matches_prefix _s4742_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4743_ _) =>
+ (match (string_drop _s4742_ _s4743_) with
+ | _s4744_ =>
+ (sep_matches_prefix _s4744_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4745_ _) =>
+ (match (string_drop _s4744_ _s4745_) with
+ | _s4746_ =>
+ (freg_name_matches_prefix _s4746_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4747_ _) =>
+ let p0_ :=
+ string_drop _s4746_ _s4747_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJN_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s4715_ (_s4716_ : string) : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))) :=
+ (match _s4716_ with
+ | _s4717_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s4717_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJ_D, existT _ _s4718_ _) =>
+ (match (string_drop _s4717_ _s4718_) with
+ | _s4719_ =>
+ (spc_matches_prefix _s4719_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4720_ _) =>
+ (match (string_drop _s4719_ _s4720_) with
+ | _s4721_ =>
+ (freg_name_matches_prefix _s4721_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4722_ _) =>
+ (match (string_drop _s4721_ _s4722_) with
+ | _s4723_ =>
+ (sep_matches_prefix _s4723_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4724_ _) =>
+ (match (string_drop _s4723_ _s4724_) with
+ | _s4725_ =>
+ (freg_name_matches_prefix _s4725_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4726_ _) =>
+ (match (string_drop _s4725_ _s4726_) with
+ | _s4727_ =>
+ (sep_matches_prefix _s4727_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4728_ _) =>
+ (match (string_drop _s4727_ _s4728_) with
+ | _s4729_ =>
+ (freg_name_matches_prefix _s4729_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4730_ _) =>
+ let p0_ :=
+ string_drop _s4729_ _s4730_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJ_D, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5))).
+
+Definition _s4698_ (_s4699_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4699_ with
+ | _s4700_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4700_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_S, existT _ _s4701_ _) =>
+ (match (string_drop _s4700_ _s4701_) with
+ | _s4702_ =>
+ (spc_matches_prefix _s4702_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4703_ _) =>
+ (match (string_drop _s4702_ _s4703_) with
+ | _s4704_ =>
+ (freg_name_matches_prefix _s4704_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4705_ _) =>
+ (match (string_drop _s4704_ _s4705_) with
+ | _s4706_ =>
+ (sep_matches_prefix _s4706_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4707_ _) =>
+ (match (string_drop _s4706_ _s4707_) with
+ | _s4708_ =>
+ (freg_name_matches_prefix _s4708_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4709_ _) =>
+ (match (string_drop _s4708_ _s4709_) with
+ | _s4710_ =>
+ (sep_matches_prefix _s4710_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4711_ _) =>
+ (match (string_drop _s4710_ _s4711_) with
+ | _s4712_ =>
+ (frm_mnemonic_matches_prefix _s4712_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4713_ _) =>
+ let p0_ :=
+ string_drop _s4712_ _s4713_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4681_ (_s4682_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4682_ with
+ | _s4683_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4683_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_D, existT _ _s4684_ _) =>
+ (match (string_drop _s4683_ _s4684_) with
+ | _s4685_ =>
+ (spc_matches_prefix _s4685_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4686_ _) =>
+ (match (string_drop _s4685_ _s4686_) with
+ | _s4687_ =>
+ (freg_name_matches_prefix _s4687_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4688_ _) =>
+ (match (string_drop _s4687_ _s4688_) with
+ | _s4689_ =>
+ (sep_matches_prefix _s4689_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4690_ _) =>
+ (match (string_drop _s4689_ _s4690_) with
+ | _s4691_ =>
+ (freg_name_matches_prefix _s4691_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4692_ _) =>
+ (match (string_drop _s4691_ _s4692_) with
+ | _s4693_ =>
+ (sep_matches_prefix _s4693_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4694_ _) =>
+ (match (string_drop _s4693_ _s4694_) with
+ | _s4695_ =>
+ (frm_mnemonic_matches_prefix _s4695_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4696_ _) =>
+ let p0_ :=
+ string_drop _s4695_ _s4696_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4664_ (_s4665_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4665_ with
+ | _s4666_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4666_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_LU, existT _ _s4667_ _) =>
+ (match (string_drop _s4666_ _s4667_) with
+ | _s4668_ =>
+ (spc_matches_prefix _s4668_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4669_ _) =>
+ (match (string_drop _s4668_ _s4669_) with
+ | _s4670_ =>
+ (freg_name_matches_prefix _s4670_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4671_ _) =>
+ (match (string_drop _s4670_ _s4671_) with
+ | _s4672_ =>
+ (sep_matches_prefix _s4672_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4673_ _) =>
+ (match (string_drop _s4672_ _s4673_) with
+ | _s4674_ =>
+ (reg_name_matches_prefix _s4674_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4675_ _) =>
+ (match (string_drop _s4674_ _s4675_) with
+ | _s4676_ =>
+ (sep_matches_prefix _s4676_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4677_ _) =>
+ (match (string_drop _s4676_ _s4677_) with
+ | _s4678_ =>
+ (frm_mnemonic_matches_prefix _s4678_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4679_ _) =>
+ let p0_ :=
+ string_drop _s4678_ _s4679_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_LU, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4647_ (_s4648_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4648_ with
+ | _s4649_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4649_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_L, existT _ _s4650_ _) =>
+ (match (string_drop _s4649_ _s4650_) with
+ | _s4651_ =>
+ (spc_matches_prefix _s4651_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4652_ _) =>
+ (match (string_drop _s4651_ _s4652_) with
+ | _s4653_ =>
+ (freg_name_matches_prefix _s4653_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4654_ _) =>
+ (match (string_drop _s4653_ _s4654_) with
+ | _s4655_ =>
+ (sep_matches_prefix _s4655_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4656_ _) =>
+ (match (string_drop _s4655_ _s4656_) with
+ | _s4657_ =>
+ (reg_name_matches_prefix _s4657_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4658_ _) =>
+ (match (string_drop _s4657_ _s4658_) with
+ | _s4659_ =>
+ (sep_matches_prefix _s4659_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4660_ _) =>
+ (match (string_drop _s4659_ _s4660_) with
+ | _s4661_ =>
+ (frm_mnemonic_matches_prefix _s4661_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4662_ _) =>
+ let p0_ :=
+ string_drop _s4661_ _s4662_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_L, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4630_ (_s4631_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4631_ with
+ | _s4632_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4632_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_LU_D, existT _ _s4633_ _) =>
+ (match (string_drop _s4632_ _s4633_) with
+ | _s4634_ =>
+ (spc_matches_prefix _s4634_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4635_ _) =>
+ (match (string_drop _s4634_ _s4635_) with
+ | _s4636_ =>
+ (reg_name_matches_prefix _s4636_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4637_ _) =>
+ (match (string_drop _s4636_ _s4637_) with
+ | _s4638_ =>
+ (sep_matches_prefix _s4638_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4639_ _) =>
+ (match (string_drop _s4638_ _s4639_) with
+ | _s4640_ =>
+ (freg_name_matches_prefix _s4640_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4641_ _) =>
+ (match (string_drop _s4640_ _s4641_) with
+ | _s4642_ =>
+ (sep_matches_prefix _s4642_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4643_ _) =>
+ (match (string_drop _s4642_ _s4643_) with
+ | _s4644_ =>
+ (frm_mnemonic_matches_prefix _s4644_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4645_ _) =>
+ let p0_ :=
+ string_drop _s4644_ _s4645_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_LU_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4613_ (_s4614_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4614_ with
+ | _s4615_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4615_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_L_D, existT _ _s4616_ _) =>
+ (match (string_drop _s4615_ _s4616_) with
+ | _s4617_ =>
+ (spc_matches_prefix _s4617_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4618_ _) =>
+ (match (string_drop _s4617_ _s4618_) with
+ | _s4619_ =>
+ (reg_name_matches_prefix _s4619_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4620_ _) =>
+ (match (string_drop _s4619_ _s4620_) with
+ | _s4621_ =>
+ (sep_matches_prefix _s4621_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4622_ _) =>
+ (match (string_drop _s4621_ _s4622_) with
+ | _s4623_ =>
+ (freg_name_matches_prefix _s4623_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4624_ _) =>
+ (match (string_drop _s4623_ _s4624_) with
+ | _s4625_ =>
+ (sep_matches_prefix _s4625_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4626_ _) =>
+ (match (string_drop _s4625_ _s4626_) with
+ | _s4627_ =>
+ (frm_mnemonic_matches_prefix _s4627_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4628_ _) =>
+ let p0_ :=
+ string_drop _s4627_ _s4628_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_L_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4596_ (_s4597_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4597_ with
+ | _s4598_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4598_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_WU, existT _ _s4599_ _) =>
+ (match (string_drop _s4598_ _s4599_) with
+ | _s4600_ =>
+ (spc_matches_prefix _s4600_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4601_ _) =>
+ (match (string_drop _s4600_ _s4601_) with
+ | _s4602_ =>
+ (freg_name_matches_prefix _s4602_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4603_ _) =>
+ (match (string_drop _s4602_ _s4603_) with
+ | _s4604_ =>
+ (sep_matches_prefix _s4604_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4605_ _) =>
+ (match (string_drop _s4604_ _s4605_) with
+ | _s4606_ =>
+ (reg_name_matches_prefix _s4606_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4607_ _) =>
+ (match (string_drop _s4606_ _s4607_) with
+ | _s4608_ =>
+ (sep_matches_prefix _s4608_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4609_ _) =>
+ (match (string_drop _s4608_ _s4609_) with
+ | _s4610_ =>
+ (frm_mnemonic_matches_prefix _s4610_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4611_ _) =>
+ let p0_ :=
+ string_drop _s4610_ _s4611_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_WU, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4579_ (_s4580_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4580_ with
+ | _s4581_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4581_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_W, existT _ _s4582_ _) =>
+ (match (string_drop _s4581_ _s4582_) with
+ | _s4583_ =>
+ (spc_matches_prefix _s4583_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4584_ _) =>
+ (match (string_drop _s4583_ _s4584_) with
+ | _s4585_ =>
+ (freg_name_matches_prefix _s4585_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4586_ _) =>
+ (match (string_drop _s4585_ _s4586_) with
+ | _s4587_ =>
+ (sep_matches_prefix _s4587_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4588_ _) =>
+ (match (string_drop _s4587_ _s4588_) with
+ | _s4589_ =>
+ (reg_name_matches_prefix _s4589_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4590_ _) =>
+ (match (string_drop _s4589_ _s4590_) with
+ | _s4591_ =>
+ (sep_matches_prefix _s4591_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4592_ _) =>
+ (match (string_drop _s4591_ _s4592_) with
+ | _s4593_ =>
+ (frm_mnemonic_matches_prefix _s4593_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4594_ _) =>
+ let p0_ :=
+ string_drop _s4593_ _s4594_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_D_W, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4562_ (_s4563_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4563_ with
+ | _s4564_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4564_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_WU_D, existT _ _s4565_ _) =>
+ (match (string_drop _s4564_ _s4565_) with
+ | _s4566_ =>
+ (spc_matches_prefix _s4566_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4567_ _) =>
+ (match (string_drop _s4566_ _s4567_) with
+ | _s4568_ =>
+ (reg_name_matches_prefix _s4568_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4569_ _) =>
+ (match (string_drop _s4568_ _s4569_) with
+ | _s4570_ =>
+ (sep_matches_prefix _s4570_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4571_ _) =>
+ (match (string_drop _s4570_ _s4571_) with
+ | _s4572_ =>
+ (freg_name_matches_prefix _s4572_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4573_ _) =>
+ (match (string_drop _s4572_ _s4573_) with
+ | _s4574_ =>
+ (sep_matches_prefix _s4574_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4575_ _) =>
+ (match (string_drop _s4574_ _s4575_) with
+ | _s4576_ =>
+ (frm_mnemonic_matches_prefix _s4576_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4577_ _) =>
+ let p0_ :=
+ string_drop _s4576_ _s4577_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_WU_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4545_ (_s4546_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4546_ with
+ | _s4547_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4547_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_W_D, existT _ _s4548_ _) =>
+ (match (string_drop _s4547_ _s4548_) with
+ | _s4549_ =>
+ (spc_matches_prefix _s4549_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4550_ _) =>
+ (match (string_drop _s4549_ _s4550_) with
+ | _s4551_ =>
+ (reg_name_matches_prefix _s4551_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4552_ _) =>
+ (match (string_drop _s4551_ _s4552_) with
+ | _s4553_ =>
+ (sep_matches_prefix _s4553_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4554_ _) =>
+ (match (string_drop _s4553_ _s4554_) with
+ | _s4555_ =>
+ (freg_name_matches_prefix _s4555_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4556_ _) =>
+ (match (string_drop _s4555_ _s4556_) with
+ | _s4557_ =>
+ (sep_matches_prefix _s4557_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4558_ _) =>
+ (match (string_drop _s4557_ _s4558_) with
+ | _s4559_ =>
+ (frm_mnemonic_matches_prefix _s4559_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4560_ _) =>
+ let p0_ :=
+ string_drop _s4559_ _s4560_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_W_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4528_ (_s4529_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4529_ with
+ | _s4530_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s4530_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSQRT_D, existT _ _s4531_ _) =>
+ (match (string_drop _s4530_ _s4531_) with
+ | _s4532_ =>
+ (spc_matches_prefix _s4532_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4533_ _) =>
+ (match (string_drop _s4532_ _s4533_) with
+ | _s4534_ =>
+ (freg_name_matches_prefix _s4534_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4535_ _) =>
+ (match (string_drop _s4534_ _s4535_) with
+ | _s4536_ =>
+ (sep_matches_prefix _s4536_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4537_ _) =>
+ (match (string_drop _s4536_ _s4537_) with
+ | _s4538_ =>
+ (freg_name_matches_prefix _s4538_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4539_ _) =>
+ (match (string_drop _s4538_ _s4539_) with
+ | _s4540_ =>
+ (sep_matches_prefix _s4540_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4541_ _) =>
+ (match (string_drop _s4540_ _s4541_) with
+ | _s4542_ =>
+ (frm_mnemonic_matches_prefix _s4542_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4543_ _) =>
+ let p0_ :=
+ string_drop _s4542_ _s4543_ in
+ if generic_eq p0_ "" then
+ Some (FSQRT_D, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12))).
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode))).
-Definition _s2506_ (_s2507_ : string)
-: M (option ((csrop * mword 5 * mword 5 * mword 12))) :=
-
- (match _s2507_ with
- | _s2508_ =>
- (csr_mnemonic_matches_prefix _s2508_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=
+Definition _s4507_ (_s4508_ : string)
+: M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4508_ with
+ | _s4509_ =>
+ (f_bin_rm_type_mnemonic_D_matches_prefix _s4509_) >>= fun w__0 : option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2509_ _)) =>
- let _s2510_ := string_drop _s2508_ _s2509_ in
- (if ((string_startswith _s2510_ "i")) then
- (match (string_drop _s2510_ (projT1 (string_length "i"))) with
- | _s2511_ =>
- (spc_matches_prefix _s2511_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s4510_ _) =>
+ (match (string_drop _s4509_ _s4510_) with
+ | _s4511_ =>
+ (spc_matches_prefix _s4511_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4512_ _) =>
+ (match (string_drop _s4511_ _s4512_) with
+ | _s4513_ =>
+ (freg_name_matches_prefix _s4513_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4514_ _) =>
+ (match (string_drop _s4513_ _s4514_) with
+ | _s4515_ =>
+ (sep_matches_prefix _s4515_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4516_ _) =>
+ (match (string_drop _s4515_ _s4516_) with
+ | _s4517_ =>
+ (freg_name_matches_prefix _s4517_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4518_ _) =>
+ (match (string_drop _s4517_ _s4518_) with
+ | _s4519_ =>
+ (sep_matches_prefix _s4519_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4520_ _) =>
+ (match (string_drop _s4519_ _s4520_) with
+ | _s4521_ =>
+ (freg_name_matches_prefix _s4521_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s4522_ _) =>
+ (match (string_drop _s4521_ _s4522_) with
+ | _s4523_ =>
+ (sep_matches_prefix _s4523_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s4524_ _) =>
+ (match (string_drop _s4523_
+ _s4524_) with
+ | _s4525_ =>
+ (frm_mnemonic_matches_prefix
+ _s4525_) >>= fun w__8 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__8 with
+ | Some
+ (rm, existT _ _s4526_ _) =>
+ let p0_ :=
+ string_drop
+ _s4525_
+ _s4526_ in
+ if generic_eq p0_
+ "" then
+ Some
+ (op, rd, rs1, rs2, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4482_ (_s4483_ : string)
+: M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4483_ with
+ | _s4484_ =>
+ (f_madd_type_mnemonic_D_matches_prefix _s4484_) >>= fun w__0 : option ((f_madd_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s4485_ _) =>
+ (match (string_drop _s4484_ _s4485_) with
+ | _s4486_ =>
+ (spc_matches_prefix _s4486_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4487_ _) =>
+ (match (string_drop _s4486_ _s4487_) with
+ | _s4488_ =>
+ (freg_name_matches_prefix _s4488_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4489_ _) =>
+ (match (string_drop _s4488_ _s4489_) with
+ | _s4490_ =>
+ (sep_matches_prefix _s4490_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4491_ _) =>
+ (match (string_drop _s4490_ _s4491_) with
+ | _s4492_ =>
+ (freg_name_matches_prefix _s4492_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4493_ _) =>
+ (match (string_drop _s4492_ _s4493_) with
+ | _s4494_ =>
+ (sep_matches_prefix _s4494_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4495_ _) =>
+ (match (string_drop _s4494_ _s4495_) with
+ | _s4496_ =>
+ (freg_name_matches_prefix _s4496_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s4497_ _) =>
+ (match (string_drop _s4496_ _s4497_) with
+ | _s4498_ =>
+ (sep_matches_prefix _s4498_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s4499_ _) =>
+ (match (string_drop _s4498_
+ _s4499_) with
+ | _s4500_ =>
+ (freg_name_matches_prefix
+ _s4500_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__8 with
+ | Some
+ (rs3, existT _ _s4501_ _) =>
+ (match (string_drop
+ _s4500_ _s4501_) with
+ | _s4502_ =>
+ (sep_matches_prefix
+ _s4502_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__9 with
+ | Some
+ (tt, existT _ _s4503_ _) =>
+ (match (string_drop
+ _s4502_
+ _s4503_) with
+ | _s4504_ =>
+ (frm_mnemonic_matches_prefix
+ _s4504_) >>= fun w__10 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__10 with
+ | Some
+ (rm, existT _ _s4505_ _) =>
+ let p0_ :=
+ string_drop
+ _s4504_
+ _s4505_ in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (op, rd, rs1, rs2, rs3, rm)
+ else
+ None
+ | _ =>
+ None
+ end)
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4466_ (_s4467_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s4468_ := _s4467_ in
+ (if string_startswith _s4468_ "c.fsw" then
+ (match (string_drop _s4468_ (projT1 (string_length "c.fsw"))) with
+ | _s4469_ =>
+ (spc_matches_prefix _s4469_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s4470_ _) =>
+ (match (string_drop _s4469_ _s4470_) with
+ | _s4471_ =>
+ (creg_name_matches_prefix _s4471_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2512_ _)) =>
- (match (string_drop _s2511_ _s2512_) with
- | _s2513_ =>
- (reg_name_matches_prefix _s2513_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (rsc1, existT _ _s4472_ _) =>
+ (match (string_drop _s4471_ _s4472_) with
+ | _s4473_ =>
+ (sep_matches_prefix _s4473_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2514_ _)) =>
- (match (string_drop _s2513_ _s2514_) with
- | _s2515_ =>
- (sep_matches_prefix _s2515_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s4474_ _) =>
+ (match (string_drop _s4473_ _s4474_) with
+ | _s4475_ =>
+ (creg_name_matches_prefix _s4475_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2516_ _)) =>
- (match (string_drop _s2515_ _s2516_) with
- | _s2517_ =>
- (match (hex_bits_5_matches_prefix _s2517_) with
- | Some ((rs1, existT _ _s2518_ _)) =>
- (match (string_drop _s2517_ _s2518_) with
- | _s2519_ =>
- (sep_matches_prefix _s2519_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc2, existT _ _s4476_ _) =>
+ (match (string_drop _s4475_ _s4476_) with
+ | _s4477_ =>
+ (sep_matches_prefix _s4477_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s4478_ _) =>
+ match (string_drop _s4477_ _s4478_) with
+ | _s4479_ =>
+ match (hex_bits_7_matches_prefix _s4479_) with
+ | Some (v__1362, existT _ _s4480_ _) =>
+ if eq_vec (subrange_vec_dec v__1362 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1362 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1362 6 2 in
+ let p0_ := string_drop _s4479_ _s4480_ in
+ if generic_eq p0_ "" then
+ Some (rsc1, rsc2, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5))).
+
+Definition _s4450_ (_s4451_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s4452_ := _s4451_ in
+ (if string_startswith _s4452_ "c.flw" then
+ (match (string_drop _s4452_ (projT1 (string_length "c.flw"))) with
+ | _s4453_ =>
+ (spc_matches_prefix _s4453_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s4454_ _) =>
+ (match (string_drop _s4453_ _s4454_) with
+ | _s4455_ =>
+ (creg_name_matches_prefix _s4455_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rdc, existT _ _s4456_ _) =>
+ (match (string_drop _s4455_ _s4456_) with
+ | _s4457_ =>
+ (sep_matches_prefix _s4457_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s4458_ _) =>
+ (match (string_drop _s4457_ _s4458_) with
+ | _s4459_ =>
+ (creg_name_matches_prefix _s4459_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc, existT _ _s4460_ _) =>
+ (match (string_drop _s4459_ _s4460_) with
+ | _s4461_ =>
+ (sep_matches_prefix _s4461_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s4462_ _) =>
+ match (string_drop _s4461_ _s4462_) with
+ | _s4463_ =>
+ match (hex_bits_7_matches_prefix _s4463_) with
+ | Some (v__1364, existT _ _s4464_ _) =>
+ if eq_vec (subrange_vec_dec v__1364 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1364 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1364 6 2 in
+ let p0_ := string_drop _s4463_ _s4464_ in
+ if generic_eq p0_ "" then
+ Some (rdc, rsc, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5))).
+
+Definition _s4438_ (_s4439_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s4440_ := _s4439_ in
+ (if string_startswith _s4440_ "c.fswsp" then
+ (match (string_drop _s4440_ (projT1 (string_length "c.fswsp"))) with
+ | _s4441_ =>
+ (spc_matches_prefix _s4441_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s4442_ _) =>
+ (match (string_drop _s4441_ _s4442_) with
+ | _s4443_ =>
+ (reg_name_matches_prefix _s4443_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s4444_ _) =>
+ (match (string_drop _s4443_ _s4444_) with
+ | _s4445_ =>
+ (sep_matches_prefix _s4445_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s4446_ _) =>
+ match (string_drop _s4445_ _s4446_) with
+ | _s4447_ =>
+ match (hex_bits_6_matches_prefix _s4447_) with
+ | Some (uimm, existT _ _s4448_ _) =>
+ let p0_ := string_drop _s4447_ _s4448_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6))).
+
+Definition _s4426_ (_s4427_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s4428_ := _s4427_ in
+ (if string_startswith _s4428_ "c.flwsp" then
+ (match (string_drop _s4428_ (projT1 (string_length "c.flwsp"))) with
+ | _s4429_ =>
+ (spc_matches_prefix _s4429_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s4430_ _) =>
+ (match (string_drop _s4429_ _s4430_) with
+ | _s4431_ =>
+ (reg_name_matches_prefix _s4431_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s4432_ _) =>
+ (match (string_drop _s4431_ _s4432_) with
+ | _s4433_ =>
+ (sep_matches_prefix _s4433_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s4434_ _) =>
+ match (string_drop _s4433_ _s4434_) with
+ | _s4435_ =>
+ match (hex_bits_6_matches_prefix _s4435_) with
+ | Some (imm, existT _ _s4436_ _) =>
+ let p0_ := string_drop _s4435_ _s4436_ in
+ if generic_eq p0_ "" then Some (rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6)))
+ end)
+ : M (option ((mword 5 * mword 6)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6))).
+
+Definition _s4413_ (_s4414_ : string) : M (option ((f_un_op_S * mword 5 * mword 5))) :=
+ (match _s4414_ with
+ | _s4415_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s4415_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCLASS_S, existT _ _s4416_ _) =>
+ (match (string_drop _s4415_ _s4416_) with
+ | _s4417_ =>
+ (spc_matches_prefix _s4417_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4418_ _) =>
+ (match (string_drop _s4417_ _s4418_) with
+ | _s4419_ =>
+ (reg_name_matches_prefix _s4419_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4420_ _) =>
+ (match (string_drop _s4419_ _s4420_) with
+ | _s4421_ =>
+ (sep_matches_prefix _s4421_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4422_ _) =>
+ (match (string_drop _s4421_ _s4422_) with
+ | _s4423_ =>
+ (freg_name_matches_prefix _s4423_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s4424_ _) =>
+ let p0_ := string_drop _s4423_ _s4424_ in
+ if generic_eq p0_ "" then
+ Some (FCLASS_S, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5))).
+
+Definition _s4400_ (_s4401_ : string) : M (option ((f_un_op_S * mword 5 * mword 5))) :=
+ (match _s4401_ with
+ | _s4402_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s4402_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_W_X, existT _ _s4403_ _) =>
+ (match (string_drop _s4402_ _s4403_) with
+ | _s4404_ =>
+ (spc_matches_prefix _s4404_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4405_ _) =>
+ (match (string_drop _s4404_ _s4405_) with
+ | _s4406_ =>
+ (freg_name_matches_prefix _s4406_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4407_ _) =>
+ (match (string_drop _s4406_ _s4407_) with
+ | _s4408_ =>
+ (sep_matches_prefix _s4408_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4409_ _) =>
+ (match (string_drop _s4408_ _s4409_) with
+ | _s4410_ =>
+ (reg_name_matches_prefix _s4410_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s4411_ _) =>
+ let p0_ := string_drop _s4410_ _s4411_ in
+ if generic_eq p0_ "" then Some (FMV_W_X, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5))).
+
+Definition _s4387_ (_s4388_ : string) : M (option ((f_un_op_S * mword 5 * mword 5))) :=
+ (match _s4388_ with
+ | _s4389_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s4389_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_X_W, existT _ _s4390_ _) =>
+ (match (string_drop _s4389_ _s4390_) with
+ | _s4391_ =>
+ (spc_matches_prefix _s4391_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4392_ _) =>
+ (match (string_drop _s4391_ _s4392_) with
+ | _s4393_ =>
+ (reg_name_matches_prefix _s4393_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4394_ _) =>
+ (match (string_drop _s4393_ _s4394_) with
+ | _s4395_ =>
+ (sep_matches_prefix _s4395_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4396_ _) =>
+ (match (string_drop _s4395_ _s4396_) with
+ | _s4397_ =>
+ (freg_name_matches_prefix _s4397_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s4398_ _) =>
+ let p0_ := string_drop _s4397_ _s4398_ in
+ if generic_eq p0_ "" then Some (FMV_X_W, rd, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5))).
+
+Definition _s4370_ (_s4371_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s4371_ with
+ | _s4372_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s4372_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLE_S, existT _ _s4373_ _) =>
+ (match (string_drop _s4372_ _s4373_) with
+ | _s4374_ =>
+ (spc_matches_prefix _s4374_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4375_ _) =>
+ (match (string_drop _s4374_ _s4375_) with
+ | _s4376_ =>
+ (reg_name_matches_prefix _s4376_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4377_ _) =>
+ (match (string_drop _s4376_ _s4377_) with
+ | _s4378_ =>
+ (sep_matches_prefix _s4378_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4379_ _) =>
+ (match (string_drop _s4378_ _s4379_) with
+ | _s4380_ =>
+ (freg_name_matches_prefix _s4380_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4381_ _) =>
+ (match (string_drop _s4380_ _s4381_) with
+ | _s4382_ =>
+ (sep_matches_prefix _s4382_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4383_ _) =>
+ (match (string_drop _s4382_ _s4383_) with
+ | _s4384_ =>
+ (freg_name_matches_prefix _s4384_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4385_ _) =>
+ let p0_ :=
+ string_drop _s4384_ _s4385_ in
+ if generic_eq p0_ "" then
+ Some (FLE_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s4353_ (_s4354_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s4354_ with
+ | _s4355_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s4355_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLT_S, existT _ _s4356_ _) =>
+ (match (string_drop _s4355_ _s4356_) with
+ | _s4357_ =>
+ (spc_matches_prefix _s4357_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4358_ _) =>
+ (match (string_drop _s4357_ _s4358_) with
+ | _s4359_ =>
+ (reg_name_matches_prefix _s4359_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4360_ _) =>
+ (match (string_drop _s4359_ _s4360_) with
+ | _s4361_ =>
+ (sep_matches_prefix _s4361_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4362_ _) =>
+ (match (string_drop _s4361_ _s4362_) with
+ | _s4363_ =>
+ (freg_name_matches_prefix _s4363_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4364_ _) =>
+ (match (string_drop _s4363_ _s4364_) with
+ | _s4365_ =>
+ (sep_matches_prefix _s4365_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4366_ _) =>
+ (match (string_drop _s4365_ _s4366_) with
+ | _s4367_ =>
+ (freg_name_matches_prefix _s4367_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4368_ _) =>
+ let p0_ :=
+ string_drop _s4367_ _s4368_ in
+ if generic_eq p0_ "" then
+ Some (FLT_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s4336_ (_s4337_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s4337_ with
+ | _s4338_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s4338_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FEQ_S, existT _ _s4339_ _) =>
+ (match (string_drop _s4338_ _s4339_) with
+ | _s4340_ =>
+ (spc_matches_prefix _s4340_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4341_ _) =>
+ (match (string_drop _s4340_ _s4341_) with
+ | _s4342_ =>
+ (reg_name_matches_prefix _s4342_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4343_ _) =>
+ (match (string_drop _s4342_ _s4343_) with
+ | _s4344_ =>
+ (sep_matches_prefix _s4344_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4345_ _) =>
+ (match (string_drop _s4344_ _s4345_) with
+ | _s4346_ =>
+ (freg_name_matches_prefix _s4346_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4347_ _) =>
+ (match (string_drop _s4346_ _s4347_) with
+ | _s4348_ =>
+ (sep_matches_prefix _s4348_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4349_ _) =>
+ (match (string_drop _s4348_ _s4349_) with
+ | _s4350_ =>
+ (freg_name_matches_prefix _s4350_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4351_ _) =>
+ let p0_ :=
+ string_drop _s4350_ _s4351_ in
+ if generic_eq p0_ "" then
+ Some (FEQ_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s4319_ (_s4320_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s4320_ with
+ | _s4321_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s4321_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMAX_S, existT _ _s4322_ _) =>
+ (match (string_drop _s4321_ _s4322_) with
+ | _s4323_ =>
+ (spc_matches_prefix _s4323_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4324_ _) =>
+ (match (string_drop _s4323_ _s4324_) with
+ | _s4325_ =>
+ (freg_name_matches_prefix _s4325_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4326_ _) =>
+ (match (string_drop _s4325_ _s4326_) with
+ | _s4327_ =>
+ (sep_matches_prefix _s4327_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4328_ _) =>
+ (match (string_drop _s4327_ _s4328_) with
+ | _s4329_ =>
+ (freg_name_matches_prefix _s4329_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4330_ _) =>
+ (match (string_drop _s4329_ _s4330_) with
+ | _s4331_ =>
+ (sep_matches_prefix _s4331_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4332_ _) =>
+ (match (string_drop _s4331_ _s4332_) with
+ | _s4333_ =>
+ (freg_name_matches_prefix _s4333_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4334_ _) =>
+ let p0_ :=
+ string_drop _s4333_ _s4334_ in
+ if generic_eq p0_ "" then
+ Some (FMAX_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s4302_ (_s4303_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s4303_ with
+ | _s4304_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s4304_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMIN_S, existT _ _s4305_ _) =>
+ (match (string_drop _s4304_ _s4305_) with
+ | _s4306_ =>
+ (spc_matches_prefix _s4306_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4307_ _) =>
+ (match (string_drop _s4306_ _s4307_) with
+ | _s4308_ =>
+ (freg_name_matches_prefix _s4308_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4309_ _) =>
+ (match (string_drop _s4308_ _s4309_) with
+ | _s4310_ =>
+ (sep_matches_prefix _s4310_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4311_ _) =>
+ (match (string_drop _s4310_ _s4311_) with
+ | _s4312_ =>
+ (freg_name_matches_prefix _s4312_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4313_ _) =>
+ (match (string_drop _s4312_ _s4313_) with
+ | _s4314_ =>
+ (sep_matches_prefix _s4314_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4315_ _) =>
+ (match (string_drop _s4314_ _s4315_) with
+ | _s4316_ =>
+ (freg_name_matches_prefix _s4316_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4317_ _) =>
+ let p0_ :=
+ string_drop _s4316_ _s4317_ in
+ if generic_eq p0_ "" then
+ Some (FMIN_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s4285_ (_s4286_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s4286_ with
+ | _s4287_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s4287_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJX_S, existT _ _s4288_ _) =>
+ (match (string_drop _s4287_ _s4288_) with
+ | _s4289_ =>
+ (spc_matches_prefix _s4289_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4290_ _) =>
+ (match (string_drop _s4289_ _s4290_) with
+ | _s4291_ =>
+ (freg_name_matches_prefix _s4291_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4292_ _) =>
+ (match (string_drop _s4291_ _s4292_) with
+ | _s4293_ =>
+ (sep_matches_prefix _s4293_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4294_ _) =>
+ (match (string_drop _s4293_ _s4294_) with
+ | _s4295_ =>
+ (freg_name_matches_prefix _s4295_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4296_ _) =>
+ (match (string_drop _s4295_ _s4296_) with
+ | _s4297_ =>
+ (sep_matches_prefix _s4297_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4298_ _) =>
+ (match (string_drop _s4297_ _s4298_) with
+ | _s4299_ =>
+ (freg_name_matches_prefix _s4299_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4300_ _) =>
+ let p0_ :=
+ string_drop _s4299_ _s4300_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJX_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s4268_ (_s4269_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s4269_ with
+ | _s4270_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s4270_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJN_S, existT _ _s4271_ _) =>
+ (match (string_drop _s4270_ _s4271_) with
+ | _s4272_ =>
+ (spc_matches_prefix _s4272_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4273_ _) =>
+ (match (string_drop _s4272_ _s4273_) with
+ | _s4274_ =>
+ (freg_name_matches_prefix _s4274_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4275_ _) =>
+ (match (string_drop _s4274_ _s4275_) with
+ | _s4276_ =>
+ (sep_matches_prefix _s4276_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4277_ _) =>
+ (match (string_drop _s4276_ _s4277_) with
+ | _s4278_ =>
+ (freg_name_matches_prefix _s4278_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4279_ _) =>
+ (match (string_drop _s4278_ _s4279_) with
+ | _s4280_ =>
+ (sep_matches_prefix _s4280_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4281_ _) =>
+ (match (string_drop _s4280_ _s4281_) with
+ | _s4282_ =>
+ (freg_name_matches_prefix _s4282_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4283_ _) =>
+ let p0_ :=
+ string_drop _s4282_ _s4283_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJN_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s4251_ (_s4252_ : string) : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))) :=
+ (match _s4252_ with
+ | _s4253_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s4253_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJ_S, existT _ _s4254_ _) =>
+ (match (string_drop _s4253_ _s4254_) with
+ | _s4255_ =>
+ (spc_matches_prefix _s4255_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4256_ _) =>
+ (match (string_drop _s4255_ _s4256_) with
+ | _s4257_ =>
+ (freg_name_matches_prefix _s4257_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4258_ _) =>
+ (match (string_drop _s4257_ _s4258_) with
+ | _s4259_ =>
+ (sep_matches_prefix _s4259_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4260_ _) =>
+ (match (string_drop _s4259_ _s4260_) with
+ | _s4261_ =>
+ (freg_name_matches_prefix _s4261_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4262_ _) =>
+ (match (string_drop _s4261_ _s4262_) with
+ | _s4263_ =>
+ (sep_matches_prefix _s4263_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4264_ _) =>
+ (match (string_drop _s4263_ _s4264_) with
+ | _s4265_ =>
+ (freg_name_matches_prefix _s4265_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s4266_ _) =>
+ let p0_ :=
+ string_drop _s4265_ _s4266_ in
+ if generic_eq p0_ "" then
+ Some (FSGNJ_S, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5))).
+
+Definition _s4234_ (_s4235_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4235_ with
+ | _s4236_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4236_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_LU, existT _ _s4237_ _) =>
+ (match (string_drop _s4236_ _s4237_) with
+ | _s4238_ =>
+ (spc_matches_prefix _s4238_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4239_ _) =>
+ (match (string_drop _s4238_ _s4239_) with
+ | _s4240_ =>
+ (freg_name_matches_prefix _s4240_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4241_ _) =>
+ (match (string_drop _s4240_ _s4241_) with
+ | _s4242_ =>
+ (sep_matches_prefix _s4242_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4243_ _) =>
+ (match (string_drop _s4242_ _s4243_) with
+ | _s4244_ =>
+ (reg_name_matches_prefix _s4244_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4245_ _) =>
+ (match (string_drop _s4244_ _s4245_) with
+ | _s4246_ =>
+ (sep_matches_prefix _s4246_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4247_ _) =>
+ (match (string_drop _s4246_ _s4247_) with
+ | _s4248_ =>
+ (frm_mnemonic_matches_prefix _s4248_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4249_ _) =>
+ let p0_ :=
+ string_drop _s4248_ _s4249_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_LU, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4217_ (_s4218_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4218_ with
+ | _s4219_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4219_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_L, existT _ _s4220_ _) =>
+ (match (string_drop _s4219_ _s4220_) with
+ | _s4221_ =>
+ (spc_matches_prefix _s4221_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4222_ _) =>
+ (match (string_drop _s4221_ _s4222_) with
+ | _s4223_ =>
+ (freg_name_matches_prefix _s4223_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4224_ _) =>
+ (match (string_drop _s4223_ _s4224_) with
+ | _s4225_ =>
+ (sep_matches_prefix _s4225_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4226_ _) =>
+ (match (string_drop _s4225_ _s4226_) with
+ | _s4227_ =>
+ (reg_name_matches_prefix _s4227_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4228_ _) =>
+ (match (string_drop _s4227_ _s4228_) with
+ | _s4229_ =>
+ (sep_matches_prefix _s4229_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4230_ _) =>
+ (match (string_drop _s4229_ _s4230_) with
+ | _s4231_ =>
+ (frm_mnemonic_matches_prefix _s4231_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4232_ _) =>
+ let p0_ :=
+ string_drop _s4231_ _s4232_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_L, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4200_ (_s4201_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4201_ with
+ | _s4202_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4202_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_LU_S, existT _ _s4203_ _) =>
+ (match (string_drop _s4202_ _s4203_) with
+ | _s4204_ =>
+ (spc_matches_prefix _s4204_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4205_ _) =>
+ (match (string_drop _s4204_ _s4205_) with
+ | _s4206_ =>
+ (reg_name_matches_prefix _s4206_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4207_ _) =>
+ (match (string_drop _s4206_ _s4207_) with
+ | _s4208_ =>
+ (sep_matches_prefix _s4208_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4209_ _) =>
+ (match (string_drop _s4208_ _s4209_) with
+ | _s4210_ =>
+ (freg_name_matches_prefix _s4210_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4211_ _) =>
+ (match (string_drop _s4210_ _s4211_) with
+ | _s4212_ =>
+ (sep_matches_prefix _s4212_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4213_ _) =>
+ (match (string_drop _s4212_ _s4213_) with
+ | _s4214_ =>
+ (frm_mnemonic_matches_prefix _s4214_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4215_ _) =>
+ let p0_ :=
+ string_drop _s4214_ _s4215_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_LU_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4183_ (_s4184_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4184_ with
+ | _s4185_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4185_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_L_S, existT _ _s4186_ _) =>
+ (match (string_drop _s4185_ _s4186_) with
+ | _s4187_ =>
+ (spc_matches_prefix _s4187_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4188_ _) =>
+ (match (string_drop _s4187_ _s4188_) with
+ | _s4189_ =>
+ (reg_name_matches_prefix _s4189_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4190_ _) =>
+ (match (string_drop _s4189_ _s4190_) with
+ | _s4191_ =>
+ (sep_matches_prefix _s4191_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4192_ _) =>
+ (match (string_drop _s4191_ _s4192_) with
+ | _s4193_ =>
+ (freg_name_matches_prefix _s4193_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4194_ _) =>
+ (match (string_drop _s4193_ _s4194_) with
+ | _s4195_ =>
+ (sep_matches_prefix _s4195_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4196_ _) =>
+ (match (string_drop _s4195_ _s4196_) with
+ | _s4197_ =>
+ (frm_mnemonic_matches_prefix _s4197_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4198_ _) =>
+ let p0_ :=
+ string_drop _s4197_ _s4198_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_L_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4166_ (_s4167_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4167_ with
+ | _s4168_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4168_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_WU, existT _ _s4169_ _) =>
+ (match (string_drop _s4168_ _s4169_) with
+ | _s4170_ =>
+ (spc_matches_prefix _s4170_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4171_ _) =>
+ (match (string_drop _s4170_ _s4171_) with
+ | _s4172_ =>
+ (freg_name_matches_prefix _s4172_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4173_ _) =>
+ (match (string_drop _s4172_ _s4173_) with
+ | _s4174_ =>
+ (sep_matches_prefix _s4174_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4175_ _) =>
+ (match (string_drop _s4174_ _s4175_) with
+ | _s4176_ =>
+ (reg_name_matches_prefix _s4176_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4177_ _) =>
+ (match (string_drop _s4176_ _s4177_) with
+ | _s4178_ =>
+ (sep_matches_prefix _s4178_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4179_ _) =>
+ (match (string_drop _s4178_ _s4179_) with
+ | _s4180_ =>
+ (frm_mnemonic_matches_prefix _s4180_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4181_ _) =>
+ let p0_ :=
+ string_drop _s4180_ _s4181_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_WU, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4149_ (_s4150_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4150_ with
+ | _s4151_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4151_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_W, existT _ _s4152_ _) =>
+ (match (string_drop _s4151_ _s4152_) with
+ | _s4153_ =>
+ (spc_matches_prefix _s4153_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4154_ _) =>
+ (match (string_drop _s4153_ _s4154_) with
+ | _s4155_ =>
+ (freg_name_matches_prefix _s4155_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4156_ _) =>
+ (match (string_drop _s4155_ _s4156_) with
+ | _s4157_ =>
+ (sep_matches_prefix _s4157_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4158_ _) =>
+ (match (string_drop _s4157_ _s4158_) with
+ | _s4159_ =>
+ (reg_name_matches_prefix _s4159_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4160_ _) =>
+ (match (string_drop _s4159_ _s4160_) with
+ | _s4161_ =>
+ (sep_matches_prefix _s4161_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4162_ _) =>
+ (match (string_drop _s4161_ _s4162_) with
+ | _s4163_ =>
+ (frm_mnemonic_matches_prefix _s4163_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4164_ _) =>
+ let p0_ :=
+ string_drop _s4163_ _s4164_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_S_W, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4132_ (_s4133_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4133_ with
+ | _s4134_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4134_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_WU_S, existT _ _s4135_ _) =>
+ (match (string_drop _s4134_ _s4135_) with
+ | _s4136_ =>
+ (spc_matches_prefix _s4136_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4137_ _) =>
+ (match (string_drop _s4136_ _s4137_) with
+ | _s4138_ =>
+ (reg_name_matches_prefix _s4138_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4139_ _) =>
+ (match (string_drop _s4138_ _s4139_) with
+ | _s4140_ =>
+ (sep_matches_prefix _s4140_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4141_ _) =>
+ (match (string_drop _s4140_ _s4141_) with
+ | _s4142_ =>
+ (freg_name_matches_prefix _s4142_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4143_ _) =>
+ (match (string_drop _s4142_ _s4143_) with
+ | _s4144_ =>
+ (sep_matches_prefix _s4144_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4145_ _) =>
+ (match (string_drop _s4144_ _s4145_) with
+ | _s4146_ =>
+ (frm_mnemonic_matches_prefix _s4146_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4147_ _) =>
+ let p0_ :=
+ string_drop _s4146_ _s4147_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_WU_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4115_ (_s4116_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4116_ with
+ | _s4117_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4117_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_W_S, existT _ _s4118_ _) =>
+ (match (string_drop _s4117_ _s4118_) with
+ | _s4119_ =>
+ (spc_matches_prefix _s4119_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4120_ _) =>
+ (match (string_drop _s4119_ _s4120_) with
+ | _s4121_ =>
+ (reg_name_matches_prefix _s4121_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4122_ _) =>
+ (match (string_drop _s4121_ _s4122_) with
+ | _s4123_ =>
+ (sep_matches_prefix _s4123_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4124_ _) =>
+ (match (string_drop _s4123_ _s4124_) with
+ | _s4125_ =>
+ (freg_name_matches_prefix _s4125_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4126_ _) =>
+ (match (string_drop _s4125_ _s4126_) with
+ | _s4127_ =>
+ (sep_matches_prefix _s4127_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4128_ _) =>
+ (match (string_drop _s4127_ _s4128_) with
+ | _s4129_ =>
+ (frm_mnemonic_matches_prefix _s4129_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4130_ _) =>
+ let p0_ :=
+ string_drop _s4129_ _s4130_ in
+ if generic_eq p0_ "" then
+ Some (FCVT_W_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4098_ (_s4099_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4099_ with
+ | _s4100_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s4100_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSQRT_S, existT _ _s4101_ _) =>
+ (match (string_drop _s4100_ _s4101_) with
+ | _s4102_ =>
+ (spc_matches_prefix _s4102_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4103_ _) =>
+ (match (string_drop _s4102_ _s4103_) with
+ | _s4104_ =>
+ (freg_name_matches_prefix _s4104_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4105_ _) =>
+ (match (string_drop _s4104_ _s4105_) with
+ | _s4106_ =>
+ (sep_matches_prefix _s4106_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4107_ _) =>
+ (match (string_drop _s4106_ _s4107_) with
+ | _s4108_ =>
+ (freg_name_matches_prefix _s4108_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4109_ _) =>
+ (match (string_drop _s4108_ _s4109_) with
+ | _s4110_ =>
+ (sep_matches_prefix _s4110_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4111_ _) =>
+ (match (string_drop _s4110_ _s4111_) with
+ | _s4112_ =>
+ (frm_mnemonic_matches_prefix _s4112_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s4113_ _) =>
+ let p0_ :=
+ string_drop _s4112_ _s4113_ in
+ if generic_eq p0_ "" then
+ Some (FSQRT_S, rd, rs1, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4077_ (_s4078_ : string)
+: M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4078_ with
+ | _s4079_ =>
+ (f_bin_rm_type_mnemonic_S_matches_prefix _s4079_) >>= fun w__0 : option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s4080_ _) =>
+ (match (string_drop _s4079_ _s4080_) with
+ | _s4081_ =>
+ (spc_matches_prefix _s4081_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4082_ _) =>
+ (match (string_drop _s4081_ _s4082_) with
+ | _s4083_ =>
+ (freg_name_matches_prefix _s4083_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4084_ _) =>
+ (match (string_drop _s4083_ _s4084_) with
+ | _s4085_ =>
+ (sep_matches_prefix _s4085_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4086_ _) =>
+ (match (string_drop _s4085_ _s4086_) with
+ | _s4087_ =>
+ (freg_name_matches_prefix _s4087_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4088_ _) =>
+ (match (string_drop _s4087_ _s4088_) with
+ | _s4089_ =>
+ (sep_matches_prefix _s4089_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4090_ _) =>
+ (match (string_drop _s4089_ _s4090_) with
+ | _s4091_ =>
+ (freg_name_matches_prefix _s4091_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s4092_ _) =>
+ (match (string_drop _s4091_ _s4092_) with
+ | _s4093_ =>
+ (sep_matches_prefix _s4093_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s4094_ _) =>
+ (match (string_drop _s4093_
+ _s4094_) with
+ | _s4095_ =>
+ (frm_mnemonic_matches_prefix
+ _s4095_) >>= fun w__8 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__8 with
+ | Some
+ (rm, existT _ _s4096_ _) =>
+ let p0_ :=
+ string_drop
+ _s4095_
+ _s4096_ in
+ if generic_eq p0_
+ "" then
+ Some
+ (op, rd, rs1, rs2, rm)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4052_ (_s4053_ : string)
+: M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode))) :=
+ (match _s4053_ with
+ | _s4054_ =>
+ (f_madd_type_mnemonic_S_matches_prefix _s4054_) >>= fun w__0 : option ((f_madd_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s4055_ _) =>
+ (match (string_drop _s4054_ _s4055_) with
+ | _s4056_ =>
+ (spc_matches_prefix _s4056_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4057_ _) =>
+ (match (string_drop _s4056_ _s4057_) with
+ | _s4058_ =>
+ (freg_name_matches_prefix _s4058_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4059_ _) =>
+ (match (string_drop _s4058_ _s4059_) with
+ | _s4060_ =>
+ (sep_matches_prefix _s4060_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4061_ _) =>
+ (match (string_drop _s4060_ _s4061_) with
+ | _s4062_ =>
+ (freg_name_matches_prefix _s4062_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s4063_ _) =>
+ (match (string_drop _s4062_ _s4063_) with
+ | _s4064_ =>
+ (sep_matches_prefix _s4064_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4065_ _) =>
+ (match (string_drop _s4064_ _s4065_) with
+ | _s4066_ =>
+ (freg_name_matches_prefix _s4066_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s4067_ _) =>
+ (match (string_drop _s4066_ _s4067_) with
+ | _s4068_ =>
+ (sep_matches_prefix _s4068_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s4069_ _) =>
+ (match (string_drop _s4068_
+ _s4069_) with
+ | _s4070_ =>
+ (freg_name_matches_prefix
+ _s4070_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__8 with
+ | Some
+ (rs3, existT _ _s4071_ _) =>
+ (match (string_drop
+ _s4070_ _s4071_) with
+ | _s4072_ =>
+ (sep_matches_prefix
+ _s4072_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__9 with
+ | Some
+ (tt, existT _ _s4073_ _) =>
+ (match (string_drop
+ _s4072_
+ _s4073_) with
+ | _s4074_ =>
+ (frm_mnemonic_matches_prefix
+ _s4074_) >>= fun w__10 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__10 with
+ | Some
+ (rm, existT _ _s4075_ _) =>
+ let p0_ :=
+ string_drop
+ _s4074_
+ _s4075_ in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (op, rd, rs1, rs2, rs3, rm)
+ else
+ None
+ | _ =>
+ None
+ end)
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode))).
+
+Definition _s4028_ (_s4029_ : string) : M (option ((word_width * mword 5 * mword 12 * mword 5))) :=
+ let _s4030_ := _s4029_ in
+ (if string_startswith _s4030_ "fs" then
+ (match (string_drop _s4030_ (projT1 (string_length "fs"))) with
+ | _s4031_ =>
+ (size_mnemonic_matches_prefix _s4031_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (width, existT _ _s4032_ _) =>
+ (match (string_drop _s4031_ _s4032_) with
+ | _s4033_ =>
+ (spc_matches_prefix _s4033_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4034_ _) =>
+ (match (string_drop _s4033_ _s4034_) with
+ | _s4035_ =>
+ (freg_name_matches_prefix _s4035_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rs2, existT _ _s4036_ _) =>
+ (match (string_drop _s4035_ _s4036_) with
+ | _s4037_ =>
+ (sep_matches_prefix _s4037_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4038_ _) =>
+ (match (string_drop _s4037_ _s4038_) with
+ | _s4039_ =>
+ (match (hex_bits_12_matches_prefix _s4039_) with
+ | Some (imm, existT _ _s4040_ _) =>
+ (match (string_drop _s4039_ _s4040_) with
+ | _s4041_ =>
+ (opt_spc_matches_prefix _s4041_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s2520_ _)) =>
- (match (string_drop _s2519_ _s2520_) with
- | _s2521_ =>
- (csr_name_map_matches_prefix _s2521_) >>= fun w__5 : option ((mword 12 * {n : Z & ArithFact (n >=
- 0)})) =>
- returnm ((match w__5 with
- | Some
- ((csr, existT _ _s2522_ _)) =>
- let p0_ :=
- string_drop _s2521_ _s2522_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, csr))
- else None
- | _ => None
- end)
- : option ((csrop * mword 5 * mword 5 * mword 12)))
- end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ | Some (tt, existT _ _s4042_ _) =>
+ let _s4043_ := string_drop _s4041_ _s4042_ in
+ (if string_startswith _s4043_ "(" then
+ (match (string_drop _s4043_
+ (projT1
+ (string_length "("))) with
+ | _s4044_ =>
+ (opt_spc_matches_prefix _s4044_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4045_ _) =>
+ (match (string_drop _s4044_ _s4045_) with
+ | _s4046_ =>
+ (reg_name_matches_prefix
+ _s4046_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some
+ (rs1, existT _ _s4047_ _) =>
+ (match (string_drop _s4046_
+ _s4047_) with
+ | _s4048_ =>
+ (opt_spc_matches_prefix
+ _s4048_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__7 with
+ | Some
+ (tt, existT _ _s4049_ _) =>
+ let _s4050_ :=
+ string_drop
+ _s4048_
+ _s4049_ in
+ if string_startswith
+ _s4050_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s4050_
+ (projT1
+ (string_length
+ ")")) in
+ if generic_eq
+ p0_
+ ""
+ then
+ Some
+ (width, rs2, imm, rs1)
+ else None
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- else returnm (None : option ((csrop * mword 5 * mword 5 * mword 12))))
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5))).
+
+Definition _s4004_ (_s4005_ : string) : M (option ((word_width * mword 5 * mword 12 * mword 5))) :=
+ let _s4006_ := _s4005_ in
+ (if string_startswith _s4006_ "fl" then
+ (match (string_drop _s4006_ (projT1 (string_length "fl"))) with
+ | _s4007_ =>
+ (size_mnemonic_matches_prefix _s4007_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (width, existT _ _s4008_ _) =>
+ (match (string_drop _s4007_ _s4008_) with
+ | _s4009_ =>
+ (spc_matches_prefix _s4009_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s4010_ _) =>
+ (match (string_drop _s4009_ _s4010_) with
+ | _s4011_ =>
+ (freg_name_matches_prefix _s4011_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s4012_ _) =>
+ (match (string_drop _s4011_ _s4012_) with
+ | _s4013_ =>
+ (sep_matches_prefix _s4013_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s4014_ _) =>
+ (match (string_drop _s4013_ _s4014_) with
+ | _s4015_ =>
+ (match (hex_bits_12_matches_prefix _s4015_) with
+ | Some (imm, existT _ _s4016_ _) =>
+ (match (string_drop _s4015_ _s4016_) with
+ | _s4017_ =>
+ (opt_spc_matches_prefix _s4017_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (tt, existT _ _s4018_ _) =>
+ let _s4019_ := string_drop _s4017_ _s4018_ in
+ (if string_startswith _s4019_ "(" then
+ (match (string_drop _s4019_
+ (projT1
+ (string_length "("))) with
+ | _s4020_ =>
+ (opt_spc_matches_prefix _s4020_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4021_ _) =>
+ (match (string_drop _s4020_ _s4021_) with
+ | _s4022_ =>
+ (reg_name_matches_prefix
+ _s4022_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some
+ (rs1, existT _ _s4023_ _) =>
+ (match (string_drop _s4022_
+ _s4023_) with
+ | _s4024_ =>
+ (opt_spc_matches_prefix
+ _s4024_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__7 with
+ | Some
+ (tt, existT _ _s4025_ _) =>
+ let _s4026_ :=
+ string_drop
+ _s4024_
+ _s4025_ in
+ if string_startswith
+ _s4026_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s4026_
+ (projT1
+ (string_length
+ ")")) in
+ if generic_eq
+ p0_
+ ""
+ then
+ Some
+ (width, rd, imm, rs1)
+ else None
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5))).
+
+Definition _s3987_ (_s3988_ : string) : M (option ((csrop * mword 5 * mword 12 * mword 5))) :=
+ (match _s3988_ with
+ | _s3989_ =>
+ (csr_mnemonic_matches_prefix _s3989_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s3990_ _) =>
+ (match (string_drop _s3989_ _s3990_) with
+ | _s3991_ =>
+ (spc_matches_prefix _s3991_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s3992_ _) =>
+ (match (string_drop _s3991_ _s3992_) with
+ | _s3993_ =>
+ (reg_name_matches_prefix _s3993_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s3994_ _) =>
+ (match (string_drop _s3993_ _s3994_) with
+ | _s3995_ =>
+ (sep_matches_prefix _s3995_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s3996_ _) =>
+ (match (string_drop _s3995_ _s3996_) with
+ | _s3997_ =>
+ (csr_name_map_matches_prefix _s3997_) >>= fun w__4 : option ((mword 12 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (csr, existT _ _s3998_ _) =>
+ (match (string_drop _s3997_ _s3998_) with
+ | _s3999_ =>
+ (sep_matches_prefix _s3999_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s4000_ _) =>
+ (match (string_drop _s3999_ _s4000_) with
+ | _s4001_ =>
+ (reg_name_matches_prefix _s4001_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs1, existT _ _s4002_ _) =>
+ let p0_ :=
+ string_drop _s4001_ _s4002_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, csr, rs1)
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12)))
- end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12))).
-
-Definition _s2487_ (_s2488_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s2489_ := _s2488_ in
- (if ((string_startswith _s2489_ "rem")) then
- (match (string_drop _s2489_ (projT1 (string_length "rem"))) with
- | _s2490_ =>
- (maybe_not_u_matches_prefix _s2490_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5))).
+
+Definition _s3969_ (_s3970_ : string) : M (option ((csrop * mword 5 * mword 12 * mword 5))) :=
+ (match _s3970_ with
+ | _s3971_ =>
+ (csr_mnemonic_matches_prefix _s3971_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s3972_ _) =>
+ let _s3973_ := string_drop _s3971_ _s3972_ in
+ (if string_startswith _s3973_ "i" then
+ (match (string_drop _s3973_ (projT1 (string_length "i"))) with
+ | _s3974_ =>
+ (spc_matches_prefix _s3974_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s3975_ _) =>
+ (match (string_drop _s3974_ _s3975_) with
+ | _s3976_ =>
+ (reg_name_matches_prefix _s3976_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s3977_ _) =>
+ (match (string_drop _s3976_ _s3977_) with
+ | _s3978_ =>
+ (sep_matches_prefix _s3978_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s3979_ _) =>
+ (match (string_drop _s3978_ _s3979_) with
+ | _s3980_ =>
+ (csr_name_map_matches_prefix _s3980_) >>= fun w__4 : option ((mword 12 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (csr, existT _ _s3981_ _) =>
+ (match (string_drop _s3980_ _s3981_) with
+ | _s3982_ =>
+ (sep_matches_prefix _s3982_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__5 with
+ | Some (tt, existT _ _s3983_ _) =>
+ match (string_drop _s3982_ _s3983_) with
+ | _s3984_ =>
+ match (hex_bits_5_matches_prefix
+ _s3984_) with
+ | Some (rs1, existT _ _s3985_ _) =>
+ let p0_ :=
+ string_drop _s3984_ _s3985_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, csr, rs1)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ else returnm None)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5))).
+
+Definition _s3950_ (_s3951_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5))) :=
+ let _s3952_ := _s3951_ in
+ (if string_startswith _s3952_ "rem" then
+ (match (string_drop _s3952_ (projT1 (string_length "rem"))) with
+ | _s3953_ =>
+ (maybe_not_u_matches_prefix _s3953_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s2491_ _)) =>
- let _s2492_ := string_drop _s2490_ _s2491_ in
- (if ((string_startswith _s2492_ "w")) then
- (match (string_drop _s2492_ (projT1 (string_length "w"))) with
- | _s2493_ =>
- (spc_matches_prefix _s2493_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s3954_ _) =>
+ let _s3955_ := string_drop _s3953_ _s3954_ in
+ (if string_startswith _s3955_ "w" then
+ (match (string_drop _s3955_ (projT1 (string_length "w"))) with
+ | _s3956_ =>
+ (spc_matches_prefix _s3956_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2494_ _)) =>
- (match (string_drop _s2493_ _s2494_) with
- | _s2495_ =>
- (reg_name_matches_prefix _s2495_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3957_ _) =>
+ (match (string_drop _s3956_ _s3957_) with
+ | _s3958_ =>
+ (reg_name_matches_prefix _s3958_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2496_ _)) =>
- (match (string_drop _s2495_ _s2496_) with
- | _s2497_ =>
- (sep_matches_prefix _s2497_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3959_ _) =>
+ (match (string_drop _s3958_ _s3959_) with
+ | _s3960_ =>
+ (sep_matches_prefix _s3960_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2498_ _)) =>
- (match (string_drop _s2497_ _s2498_) with
- | _s2499_ =>
- (reg_name_matches_prefix _s2499_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3961_ _) =>
+ (match (string_drop _s3960_ _s3961_) with
+ | _s3962_ =>
+ (reg_name_matches_prefix _s3962_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2500_ _)) =>
- (match (string_drop _s2499_ _s2500_) with
- | _s2501_ =>
- (sep_matches_prefix _s2501_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3963_ _) =>
+ (match (string_drop _s3962_ _s3963_) with
+ | _s3964_ =>
+ (sep_matches_prefix _s3964_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2502_ _)) =>
- (match (string_drop _s2501_ _s2502_) with
- | _s2503_ =>
- (reg_name_matches_prefix _s2503_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3965_ _) =>
+ (match (string_drop _s3964_ _s3965_) with
+ | _s3966_ =>
+ (reg_name_matches_prefix _s3966_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s2504_ _)) =>
- let p0_ :=
- string_drop _s2503_
- _s2504_ in
- if ((generic_eq p0_ ""))
- then
- Some
- ((s, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some
+ (rs2, existT _ _s3967_ _) =>
+ let p0_ :=
+ string_drop _s3966_
+ _s3967_ in
+ if generic_eq p0_ "" then
+ Some (s, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5))).
-Definition _s2468_ (_s2469_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s2470_ := _s2469_ in
- (if ((string_startswith _s2470_ "div")) then
- (match (string_drop _s2470_ (projT1 (string_length "div"))) with
- | _s2471_ =>
- (maybe_not_u_matches_prefix _s2471_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s3931_ (_s3932_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5))) :=
+ let _s3933_ := _s3932_ in
+ (if string_startswith _s3933_ "div" then
+ (match (string_drop _s3933_ (projT1 (string_length "div"))) with
+ | _s3934_ =>
+ (maybe_not_u_matches_prefix _s3934_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s2472_ _)) =>
- let _s2473_ := string_drop _s2471_ _s2472_ in
- (if ((string_startswith _s2473_ "w")) then
- (match (string_drop _s2473_ (projT1 (string_length "w"))) with
- | _s2474_ =>
- (spc_matches_prefix _s2474_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s3935_ _) =>
+ let _s3936_ := string_drop _s3934_ _s3935_ in
+ (if string_startswith _s3936_ "w" then
+ (match (string_drop _s3936_ (projT1 (string_length "w"))) with
+ | _s3937_ =>
+ (spc_matches_prefix _s3937_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2475_ _)) =>
- (match (string_drop _s2474_ _s2475_) with
- | _s2476_ =>
- (reg_name_matches_prefix _s2476_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3938_ _) =>
+ (match (string_drop _s3937_ _s3938_) with
+ | _s3939_ =>
+ (reg_name_matches_prefix _s3939_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2477_ _)) =>
- (match (string_drop _s2476_ _s2477_) with
- | _s2478_ =>
- (sep_matches_prefix _s2478_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3940_ _) =>
+ (match (string_drop _s3939_ _s3940_) with
+ | _s3941_ =>
+ (sep_matches_prefix _s3941_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2479_ _)) =>
- (match (string_drop _s2478_ _s2479_) with
- | _s2480_ =>
- (reg_name_matches_prefix _s2480_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3942_ _) =>
+ (match (string_drop _s3941_ _s3942_) with
+ | _s3943_ =>
+ (reg_name_matches_prefix _s3943_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2481_ _)) =>
- (match (string_drop _s2480_ _s2481_) with
- | _s2482_ =>
- (sep_matches_prefix _s2482_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3944_ _) =>
+ (match (string_drop _s3943_ _s3944_) with
+ | _s3945_ =>
+ (sep_matches_prefix _s3945_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2483_ _)) =>
- (match (string_drop _s2482_ _s2483_) with
- | _s2484_ =>
- (reg_name_matches_prefix _s2484_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3946_ _) =>
+ (match (string_drop _s3945_ _s3946_) with
+ | _s3947_ =>
+ (reg_name_matches_prefix _s3947_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s2485_ _)) =>
- let p0_ :=
- string_drop _s2484_
- _s2485_ in
- if ((generic_eq p0_ ""))
- then
- Some
- ((s, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some
+ (rs2, existT _ _s3948_ _) =>
+ let p0_ :=
+ string_drop _s3947_
+ _s3948_ in
+ if generic_eq p0_ "" then
+ Some (s, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5))).
-Definition _s2452_ (_s2453_ : string)
-: M (option ((mword 5 * mword 5 * mword 5))) :=
-
- let _s2454_ := _s2453_ in
- (if ((string_startswith _s2454_ "mulw")) then
- (match (string_drop _s2454_ (projT1 (string_length "mulw"))) with
- | _s2455_ =>
- (spc_matches_prefix _s2455_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3915_ (_s3916_ : string) : M (option ((mword 5 * mword 5 * mword 5))) :=
+ let _s3917_ := _s3916_ in
+ (if string_startswith _s3917_ "mulw" then
+ (match (string_drop _s3917_ (projT1 (string_length "mulw"))) with
+ | _s3918_ =>
+ (spc_matches_prefix _s3918_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2456_ _)) =>
- (match (string_drop _s2455_ _s2456_) with
- | _s2457_ =>
- (reg_name_matches_prefix _s2457_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3919_ _) =>
+ (match (string_drop _s3918_ _s3919_) with
+ | _s3920_ =>
+ (reg_name_matches_prefix _s3920_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2458_ _)) =>
- (match (string_drop _s2457_ _s2458_) with
- | _s2459_ =>
- (sep_matches_prefix _s2459_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3921_ _) =>
+ (match (string_drop _s3920_ _s3921_) with
+ | _s3922_ =>
+ (sep_matches_prefix _s3922_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2460_ _)) =>
- (match (string_drop _s2459_ _s2460_) with
- | _s2461_ =>
- (reg_name_matches_prefix _s2461_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3923_ _) =>
+ (match (string_drop _s3922_ _s3923_) with
+ | _s3924_ =>
+ (reg_name_matches_prefix _s3924_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s2462_ _)) =>
- (match (string_drop _s2461_ _s2462_) with
- | _s2463_ =>
- (sep_matches_prefix _s2463_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3925_ _) =>
+ (match (string_drop _s3924_ _s3925_) with
+ | _s3926_ =>
+ (sep_matches_prefix _s3926_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s2464_ _)) =>
- (match (string_drop _s2463_ _s2464_) with
- | _s2465_ =>
- (reg_name_matches_prefix _s2465_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3927_ _) =>
+ (match (string_drop _s3926_ _s3927_) with
+ | _s3928_ =>
+ (reg_name_matches_prefix _s3928_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((rs2, existT _ _s2466_ _)) =>
- let p0_ :=
- string_drop _s2465_ _s2466_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 5)))
+ returnm (match w__5 with
+ | Some (rs2, existT _ _s3929_ _) =>
+ let p0_ := string_drop _s3928_ _s3929_ in
+ if generic_eq p0_ "" then
+ Some (rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 5))).
-Definition _s2434_ (_s2435_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s2436_ := _s2435_ in
- (if ((string_startswith _s2436_ "rem")) then
- (match (string_drop _s2436_ (projT1 (string_length "rem"))) with
- | _s2437_ =>
- (maybe_not_u_matches_prefix _s2437_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s3897_ (_s3898_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5))) :=
+ let _s3899_ := _s3898_ in
+ (if string_startswith _s3899_ "rem" then
+ (match (string_drop _s3899_ (projT1 (string_length "rem"))) with
+ | _s3900_ =>
+ (maybe_not_u_matches_prefix _s3900_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s2438_ _)) =>
- (match (string_drop _s2437_ _s2438_) with
- | _s2439_ =>
- (spc_matches_prefix _s2439_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s3901_ _) =>
+ (match (string_drop _s3900_ _s3901_) with
+ | _s3902_ =>
+ (spc_matches_prefix _s3902_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2440_ _)) =>
- (match (string_drop _s2439_ _s2440_) with
- | _s2441_ =>
- (reg_name_matches_prefix _s2441_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3903_ _) =>
+ (match (string_drop _s3902_ _s3903_) with
+ | _s3904_ =>
+ (reg_name_matches_prefix _s3904_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2442_ _)) =>
- (match (string_drop _s2441_ _s2442_) with
- | _s2443_ =>
- (sep_matches_prefix _s2443_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3905_ _) =>
+ (match (string_drop _s3904_ _s3905_) with
+ | _s3906_ =>
+ (sep_matches_prefix _s3906_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2444_ _)) =>
- (match (string_drop _s2443_ _s2444_) with
- | _s2445_ =>
- (reg_name_matches_prefix _s2445_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3907_ _) =>
+ (match (string_drop _s3906_ _s3907_) with
+ | _s3908_ =>
+ (reg_name_matches_prefix _s3908_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2446_ _)) =>
- (match (string_drop _s2445_ _s2446_) with
- | _s2447_ =>
- (sep_matches_prefix _s2447_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3909_ _) =>
+ (match (string_drop _s3908_ _s3909_) with
+ | _s3910_ =>
+ (sep_matches_prefix _s3910_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2448_ _)) =>
- (match (string_drop _s2447_ _s2448_) with
- | _s2449_ =>
- (reg_name_matches_prefix _s2449_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3911_ _) =>
+ (match (string_drop _s3910_ _s3911_) with
+ | _s3912_ =>
+ (reg_name_matches_prefix _s3912_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s2450_ _)) =>
- let p0_ :=
- string_drop _s2449_ _s2450_ in
- if ((generic_eq p0_ "")) then
- Some
- ((s, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s3913_ _) =>
+ let p0_ :=
+ string_drop _s3912_ _s3913_ in
+ if generic_eq p0_ "" then
+ Some (s, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5))).
-Definition _s2416_ (_s2417_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s2418_ := _s2417_ in
- (if ((string_startswith _s2418_ "div")) then
- (match (string_drop _s2418_ (projT1 (string_length "div"))) with
- | _s2419_ =>
- (maybe_not_u_matches_prefix _s2419_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s3879_ (_s3880_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5))) :=
+ let _s3881_ := _s3880_ in
+ (if string_startswith _s3881_ "div" then
+ (match (string_drop _s3881_ (projT1 (string_length "div"))) with
+ | _s3882_ =>
+ (maybe_not_u_matches_prefix _s3882_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s2420_ _)) =>
- (match (string_drop _s2419_ _s2420_) with
- | _s2421_ =>
- (spc_matches_prefix _s2421_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s3883_ _) =>
+ (match (string_drop _s3882_ _s3883_) with
+ | _s3884_ =>
+ (spc_matches_prefix _s3884_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2422_ _)) =>
- (match (string_drop _s2421_ _s2422_) with
- | _s2423_ =>
- (reg_name_matches_prefix _s2423_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3885_ _) =>
+ (match (string_drop _s3884_ _s3885_) with
+ | _s3886_ =>
+ (reg_name_matches_prefix _s3886_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2424_ _)) =>
- (match (string_drop _s2423_ _s2424_) with
- | _s2425_ =>
- (sep_matches_prefix _s2425_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3887_ _) =>
+ (match (string_drop _s3886_ _s3887_) with
+ | _s3888_ =>
+ (sep_matches_prefix _s3888_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2426_ _)) =>
- (match (string_drop _s2425_ _s2426_) with
- | _s2427_ =>
- (reg_name_matches_prefix _s2427_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3889_ _) =>
+ (match (string_drop _s3888_ _s3889_) with
+ | _s3890_ =>
+ (reg_name_matches_prefix _s3890_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2428_ _)) =>
- (match (string_drop _s2427_ _s2428_) with
- | _s2429_ =>
- (sep_matches_prefix _s2429_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3891_ _) =>
+ (match (string_drop _s3890_ _s3891_) with
+ | _s3892_ =>
+ (sep_matches_prefix _s3892_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2430_ _)) =>
- (match (string_drop _s2429_ _s2430_) with
- | _s2431_ =>
- (reg_name_matches_prefix _s2431_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3893_ _) =>
+ (match (string_drop _s3892_ _s3893_) with
+ | _s3894_ =>
+ (reg_name_matches_prefix _s3894_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s2432_ _)) =>
- let p0_ :=
- string_drop _s2431_ _s2432_ in
- if ((generic_eq p0_ "")) then
- Some
- ((s, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s3895_ _) =>
+ let p0_ :=
+ string_drop _s3894_ _s3895_ in
+ if generic_eq p0_ "" then
+ Some (s, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5))).
-Definition _s2399_ (_s2400_ : string)
+Definition _s3862_ (_s3863_ : string)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5))) :=
-
- (match _s2400_ with
- | _s2401_ =>
- (mul_mnemonic_matches_prefix _s2401_) >>= fun w__0 : option (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ (match _s3863_ with
+ | _s3864_ =>
+ (mul_mnemonic_matches_prefix _s3864_) >>= fun w__0 : option (((bool * bool * bool) * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some (((high, signed1, signed2), existT _ _s2402_ _)) =>
- (match (string_drop _s2401_ _s2402_) with
- | _s2403_ =>
- (spc_matches_prefix _s2403_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some ((high, signed1, signed2), existT _ _s3865_ _) =>
+ (match (string_drop _s3864_ _s3865_) with
+ | _s3866_ =>
+ (spc_matches_prefix _s3866_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2404_ _)) =>
- (match (string_drop _s2403_ _s2404_) with
- | _s2405_ =>
- (reg_name_matches_prefix _s2405_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3867_ _) =>
+ (match (string_drop _s3866_ _s3867_) with
+ | _s3868_ =>
+ (reg_name_matches_prefix _s3868_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2406_ _)) =>
- (match (string_drop _s2405_ _s2406_) with
- | _s2407_ =>
- (sep_matches_prefix _s2407_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3869_ _) =>
+ (match (string_drop _s3868_ _s3869_) with
+ | _s3870_ =>
+ (sep_matches_prefix _s3870_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2408_ _)) =>
- (match (string_drop _s2407_ _s2408_) with
- | _s2409_ =>
- (reg_name_matches_prefix _s2409_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3871_ _) =>
+ (match (string_drop _s3870_ _s3871_) with
+ | _s3872_ =>
+ (reg_name_matches_prefix _s3872_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2410_ _)) =>
- (match (string_drop _s2409_ _s2410_) with
- | _s2411_ =>
- (sep_matches_prefix _s2411_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3873_ _) =>
+ (match (string_drop _s3872_ _s3873_) with
+ | _s3874_ =>
+ (sep_matches_prefix _s3874_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2412_ _)) =>
- (match (string_drop _s2411_ _s2412_) with
- | _s2413_ =>
- (reg_name_matches_prefix _s2413_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3875_ _) =>
+ (match (string_drop _s3874_ _s3875_) with
+ | _s3876_ =>
+ (reg_name_matches_prefix _s3876_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s2414_ _)) =>
- let p0_ :=
- string_drop _s2413_ _s2414_ in
- if ((generic_eq p0_ "")) then
- Some
- ((high, signed1, signed2, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s3877_ _) =>
+ let p0_ :=
+ string_drop _s3876_ _s3877_ in
+ if generic_eq p0_ "" then
+ Some
+ (high, signed1, signed2, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5))).
-Definition _s2387_ (_s2388_ : string)
-: M (option ((mword 5 * mword 5))) :=
-
- let _s2389_ := _s2388_ in
- (if ((string_startswith _s2389_ "c.add")) then
- (match (string_drop _s2389_ (projT1 (string_length "c.add"))) with
- | _s2390_ =>
- (spc_matches_prefix _s2390_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3850_ (_s3851_ : string) : M (option ((mword 5 * mword 5))) :=
+ let _s3852_ := _s3851_ in
+ (if string_startswith _s3852_ "c.add" then
+ (match (string_drop _s3852_ (projT1 (string_length "c.add"))) with
+ | _s3853_ =>
+ (spc_matches_prefix _s3853_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2391_ _)) =>
- (match (string_drop _s2390_ _s2391_) with
- | _s2392_ =>
- (reg_name_matches_prefix _s2392_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3854_ _) =>
+ (match (string_drop _s3853_ _s3854_) with
+ | _s3855_ =>
+ (reg_name_matches_prefix _s3855_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2393_ _)) =>
- (match (string_drop _s2392_ _s2393_) with
- | _s2394_ =>
- (sep_matches_prefix _s2394_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3856_ _) =>
+ (match (string_drop _s3855_ _s3856_) with
+ | _s3857_ =>
+ (sep_matches_prefix _s3857_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2395_ _)) =>
- (match (string_drop _s2394_ _s2395_) with
- | _s2396_ =>
- (reg_name_matches_prefix _s2396_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3858_ _) =>
+ (match (string_drop _s3857_ _s3858_) with
+ | _s3859_ =>
+ (reg_name_matches_prefix _s3859_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2397_ _)) =>
- let p0_ := string_drop _s2396_ _s2397_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 5 * mword 5)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3860_ _) =>
+ let p0_ := string_drop _s3859_ _s3860_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- else returnm (None : option ((mword 5 * mword 5))))
+ else returnm None)
: M (option ((mword 5 * mword 5))).
-Definition _s2375_ (_s2376_ : string)
-: M (option ((mword 5 * mword 5))) :=
-
- let _s2377_ := _s2376_ in
- (if ((string_startswith _s2377_ "c.mv")) then
- (match (string_drop _s2377_ (projT1 (string_length "c.mv"))) with
- | _s2378_ =>
- (spc_matches_prefix _s2378_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3838_ (_s3839_ : string) : M (option ((mword 5 * mword 5))) :=
+ let _s3840_ := _s3839_ in
+ (if string_startswith _s3840_ "c.mv" then
+ (match (string_drop _s3840_ (projT1 (string_length "c.mv"))) with
+ | _s3841_ =>
+ (spc_matches_prefix _s3841_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2379_ _)) =>
- (match (string_drop _s2378_ _s2379_) with
- | _s2380_ =>
- (reg_name_matches_prefix _s2380_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3842_ _) =>
+ (match (string_drop _s3841_ _s3842_) with
+ | _s3843_ =>
+ (reg_name_matches_prefix _s3843_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2381_ _)) =>
- (match (string_drop _s2380_ _s2381_) with
- | _s2382_ =>
- (sep_matches_prefix _s2382_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3844_ _) =>
+ (match (string_drop _s3843_ _s3844_) with
+ | _s3845_ =>
+ (sep_matches_prefix _s3845_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2383_ _)) =>
- (match (string_drop _s2382_ _s2383_) with
- | _s2384_ =>
- (reg_name_matches_prefix _s2384_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3846_ _) =>
+ (match (string_drop _s3845_ _s3846_) with
+ | _s3847_ =>
+ (reg_name_matches_prefix _s3847_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2385_ _)) =>
- let p0_ := string_drop _s2384_ _s2385_ in
- if ((generic_eq p0_ "")) then Some ((rd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 5 * mword 5)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3848_ _) =>
+ let p0_ := string_drop _s3847_ _s3848_ in
+ if generic_eq p0_ "" then Some (rd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- else returnm (None : option ((mword 5 * mword 5))))
+ else returnm None)
: M (option ((mword 5 * mword 5))).
-Definition _s2367_ (_s2368_ : string)
-: M (option (mword 5)) :=
-
- let _s2369_ := _s2368_ in
- (if ((string_startswith _s2369_ "c.jalr")) then
- (match (string_drop _s2369_ (projT1 (string_length "c.jalr"))) with
- | _s2370_ =>
- (spc_matches_prefix _s2370_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3830_ (_s3831_ : string) : M (option (mword 5)) :=
+ let _s3832_ := _s3831_ in
+ (if string_startswith _s3832_ "c.jalr" then
+ (match (string_drop _s3832_ (projT1 (string_length "c.jalr"))) with
+ | _s3833_ =>
+ (spc_matches_prefix _s3833_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2371_ _)) =>
- (match (string_drop _s2370_ _s2371_) with
- | _s2372_ =>
- (reg_name_matches_prefix _s2372_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3834_ _) =>
+ (match (string_drop _s3833_ _s3834_) with
+ | _s3835_ =>
+ (reg_name_matches_prefix _s3835_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((rs1, existT _ _s2373_ _)) =>
- let p0_ := string_drop _s2372_ _s2373_ in
- if ((generic_eq p0_ "")) then Some (rs1)
- else None
- | _ => None
- end)
- : option (mword 5))
+ returnm (match w__1 with
+ | Some (rs1, existT _ _s3836_ _) =>
+ let p0_ := string_drop _s3835_ _s3836_ in
+ if generic_eq p0_ "" then Some rs1
+ else None
+ | _ => None
+ end)
end)
: M (option (mword 5))
- | _ => returnm (None : option (mword 5))
+ | _ => returnm None
end)
: M (option (mword 5))
end)
: M (option (mword 5))
- else returnm (None : option (mword 5)))
+ else returnm None)
: M (option (mword 5)).
-Definition _s2359_ (_s2360_ : string)
-: M (option (mword 5)) :=
-
- let _s2361_ := _s2360_ in
- (if ((string_startswith _s2361_ "c.jr")) then
- (match (string_drop _s2361_ (projT1 (string_length "c.jr"))) with
- | _s2362_ =>
- (spc_matches_prefix _s2362_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3822_ (_s3823_ : string) : M (option (mword 5)) :=
+ let _s3824_ := _s3823_ in
+ (if string_startswith _s3824_ "c.jr" then
+ (match (string_drop _s3824_ (projT1 (string_length "c.jr"))) with
+ | _s3825_ =>
+ (spc_matches_prefix _s3825_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2363_ _)) =>
- (match (string_drop _s2362_ _s2363_) with
- | _s2364_ =>
- (reg_name_matches_prefix _s2364_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3826_ _) =>
+ (match (string_drop _s3825_ _s3826_) with
+ | _s3827_ =>
+ (reg_name_matches_prefix _s3827_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((rs1, existT _ _s2365_ _)) =>
- let p0_ := string_drop _s2364_ _s2365_ in
- if ((generic_eq p0_ "")) then Some (rs1)
- else None
- | _ => None
- end)
- : option (mword 5))
+ returnm (match w__1 with
+ | Some (rs1, existT _ _s3828_ _) =>
+ let p0_ := string_drop _s3827_ _s3828_ in
+ if generic_eq p0_ "" then Some rs1
+ else None
+ | _ => None
+ end)
end)
: M (option (mword 5))
- | _ => returnm (None : option (mword 5))
+ | _ => returnm None
end)
: M (option (mword 5))
end)
: M (option (mword 5))
- else returnm (None : option (mword 5)))
+ else returnm None)
: M (option (mword 5)).
-Definition _s2347_ (_s2348_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2349_ := _s2348_ in
- (if ((string_startswith _s2349_ "c.sdsp")) then
- (match (string_drop _s2349_ (projT1 (string_length "c.sdsp"))) with
- | _s2350_ =>
- (spc_matches_prefix _s2350_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3810_ (_s3811_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3812_ := _s3811_ in
+ (if string_startswith _s3812_ "c.sdsp" then
+ (match (string_drop _s3812_ (projT1 (string_length "c.sdsp"))) with
+ | _s3813_ =>
+ (spc_matches_prefix _s3813_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2351_ _)) =>
- (match (string_drop _s2350_ _s2351_) with
- | _s2352_ =>
- (reg_name_matches_prefix _s2352_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3814_ _) =>
+ (match (string_drop _s3813_ _s3814_) with
+ | _s3815_ =>
+ (reg_name_matches_prefix _s3815_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs2, existT _ _s2353_ _)) =>
- (match (string_drop _s2352_ _s2353_) with
- | _s2354_ =>
- (sep_matches_prefix _s2354_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s3816_ _) =>
+ (match (string_drop _s3815_ _s3816_) with
+ | _s3817_ =>
+ (sep_matches_prefix _s3817_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2355_ _)) =>
- match (string_drop _s2354_ _s2355_) with
- | _s2356_ =>
- match (hex_bits_6_matches_prefix _s2356_) with
- | Some ((uimm, existT _ _s2357_ _)) =>
- let p0_ := string_drop _s2356_ _s2357_ in
- if ((generic_eq p0_ "")) then Some ((rs2, uimm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3818_ _) =>
+ match (string_drop _s3817_ _s3818_) with
+ | _s3819_ =>
+ match (hex_bits_6_matches_prefix _s3819_) with
+ | Some (uimm, existT _ _s3820_ _) =>
+ let p0_ := string_drop _s3819_ _s3820_ in
+ if generic_eq p0_ "" then Some (rs2, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2335_ (_s2336_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2337_ := _s2336_ in
- (if ((string_startswith _s2337_ "c.swsp")) then
- (match (string_drop _s2337_ (projT1 (string_length "c.swsp"))) with
- | _s2338_ =>
- (spc_matches_prefix _s2338_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3798_ (_s3799_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3800_ := _s3799_ in
+ (if string_startswith _s3800_ "c.swsp" then
+ (match (string_drop _s3800_ (projT1 (string_length "c.swsp"))) with
+ | _s3801_ =>
+ (spc_matches_prefix _s3801_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2339_ _)) =>
- (match (string_drop _s2338_ _s2339_) with
- | _s2340_ =>
- (reg_name_matches_prefix _s2340_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3802_ _) =>
+ (match (string_drop _s3801_ _s3802_) with
+ | _s3803_ =>
+ (reg_name_matches_prefix _s3803_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2341_ _)) =>
- (match (string_drop _s2340_ _s2341_) with
- | _s2342_ =>
- (sep_matches_prefix _s2342_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3804_ _) =>
+ (match (string_drop _s3803_ _s3804_) with
+ | _s3805_ =>
+ (sep_matches_prefix _s3805_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2343_ _)) =>
- match (string_drop _s2342_ _s2343_) with
- | _s2344_ =>
- match (hex_bits_6_matches_prefix _s2344_) with
- | Some ((uimm, existT _ _s2345_ _)) =>
- let p0_ := string_drop _s2344_ _s2345_ in
- if ((generic_eq p0_ "")) then Some ((rd, uimm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3806_ _) =>
+ match (string_drop _s3805_ _s3806_) with
+ | _s3807_ =>
+ match (hex_bits_6_matches_prefix _s3807_) with
+ | Some (uimm, existT _ _s3808_ _) =>
+ let p0_ := string_drop _s3807_ _s3808_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2323_ (_s2324_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2325_ := _s2324_ in
- (if ((string_startswith _s2325_ "c.ldsp")) then
- (match (string_drop _s2325_ (projT1 (string_length "c.ldsp"))) with
- | _s2326_ =>
- (spc_matches_prefix _s2326_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3786_ (_s3787_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3788_ := _s3787_ in
+ (if string_startswith _s3788_ "c.ldsp" then
+ (match (string_drop _s3788_ (projT1 (string_length "c.ldsp"))) with
+ | _s3789_ =>
+ (spc_matches_prefix _s3789_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2327_ _)) =>
- (match (string_drop _s2326_ _s2327_) with
- | _s2328_ =>
- (reg_name_matches_prefix _s2328_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3790_ _) =>
+ (match (string_drop _s3789_ _s3790_) with
+ | _s3791_ =>
+ (reg_name_matches_prefix _s3791_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2329_ _)) =>
- (match (string_drop _s2328_ _s2329_) with
- | _s2330_ =>
- (sep_matches_prefix _s2330_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3792_ _) =>
+ (match (string_drop _s3791_ _s3792_) with
+ | _s3793_ =>
+ (sep_matches_prefix _s3793_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2331_ _)) =>
- match (string_drop _s2330_ _s2331_) with
- | _s2332_ =>
- match (hex_bits_6_matches_prefix _s2332_) with
- | Some ((uimm, existT _ _s2333_ _)) =>
- let p0_ := string_drop _s2332_ _s2333_ in
- if ((generic_eq p0_ "")) then Some ((rd, uimm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3794_ _) =>
+ match (string_drop _s3793_ _s3794_) with
+ | _s3795_ =>
+ match (hex_bits_6_matches_prefix _s3795_) with
+ | Some (uimm, existT _ _s3796_ _) =>
+ let p0_ := string_drop _s3795_ _s3796_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2311_ (_s2312_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2313_ := _s2312_ in
- (if ((string_startswith _s2313_ "c.lwsp")) then
- (match (string_drop _s2313_ (projT1 (string_length "c.lwsp"))) with
- | _s2314_ =>
- (spc_matches_prefix _s2314_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3774_ (_s3775_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3776_ := _s3775_ in
+ (if string_startswith _s3776_ "c.lwsp" then
+ (match (string_drop _s3776_ (projT1 (string_length "c.lwsp"))) with
+ | _s3777_ =>
+ (spc_matches_prefix _s3777_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2315_ _)) =>
- (match (string_drop _s2314_ _s2315_) with
- | _s2316_ =>
- (reg_name_matches_prefix _s2316_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3778_ _) =>
+ (match (string_drop _s3777_ _s3778_) with
+ | _s3779_ =>
+ (reg_name_matches_prefix _s3779_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2317_ _)) =>
- (match (string_drop _s2316_ _s2317_) with
- | _s2318_ =>
- (sep_matches_prefix _s2318_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3780_ _) =>
+ (match (string_drop _s3779_ _s3780_) with
+ | _s3781_ =>
+ (sep_matches_prefix _s3781_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2319_ _)) =>
- match (string_drop _s2318_ _s2319_) with
- | _s2320_ =>
- match (hex_bits_6_matches_prefix _s2320_) with
- | Some ((uimm, existT _ _s2321_ _)) =>
- let p0_ := string_drop _s2320_ _s2321_ in
- if ((generic_eq p0_ "")) then Some ((rd, uimm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3782_ _) =>
+ match (string_drop _s3781_ _s3782_) with
+ | _s3783_ =>
+ match (hex_bits_6_matches_prefix _s3783_) with
+ | Some (uimm, existT _ _s3784_ _) =>
+ let p0_ := string_drop _s3783_ _s3784_ in
+ if generic_eq p0_ "" then Some (rd, uimm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2299_ (_s2300_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2301_ := _s2300_ in
- (if ((string_startswith _s2301_ "c.slli")) then
- (match (string_drop _s2301_ (projT1 (string_length "c.slli"))) with
- | _s2302_ =>
- (spc_matches_prefix _s2302_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3762_ (_s3763_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3764_ := _s3763_ in
+ (if string_startswith _s3764_ "c.slli" then
+ (match (string_drop _s3764_ (projT1 (string_length "c.slli"))) with
+ | _s3765_ =>
+ (spc_matches_prefix _s3765_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2303_ _)) =>
- (match (string_drop _s2302_ _s2303_) with
- | _s2304_ =>
- (reg_name_matches_prefix _s2304_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3766_ _) =>
+ (match (string_drop _s3765_ _s3766_) with
+ | _s3767_ =>
+ (reg_name_matches_prefix _s3767_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2305_ _)) =>
- (match (string_drop _s2304_ _s2305_) with
- | _s2306_ =>
- (sep_matches_prefix _s2306_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3768_ _) =>
+ (match (string_drop _s3767_ _s3768_) with
+ | _s3769_ =>
+ (sep_matches_prefix _s3769_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2307_ _)) =>
- match (string_drop _s2306_ _s2307_) with
- | _s2308_ =>
- match (hex_bits_6_matches_prefix _s2308_) with
- | Some ((shamt, existT _ _s2309_ _)) =>
- let p0_ := string_drop _s2308_ _s2309_ in
- if ((generic_eq p0_ "")) then Some ((rsd, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3770_ _) =>
+ match (string_drop _s3769_ _s3770_) with
+ | _s3771_ =>
+ match (hex_bits_6_matches_prefix _s3771_) with
+ | Some (shamt, existT _ _s3772_ _) =>
+ let p0_ := string_drop _s3771_ _s3772_ in
+ if generic_eq p0_ "" then Some (rsd, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2287_ (_s2288_ : string)
-: M (option ((mword 3 * mword 8))) :=
-
- let _s2289_ := _s2288_ in
- (if ((string_startswith _s2289_ "c.bnez")) then
- (match (string_drop _s2289_ (projT1 (string_length "c.bnez"))) with
- | _s2290_ =>
- (spc_matches_prefix _s2290_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3750_ (_s3751_ : string) : M (option ((mword 3 * mword 8))) :=
+ let _s3752_ := _s3751_ in
+ (if string_startswith _s3752_ "c.bnez" then
+ (match (string_drop _s3752_ (projT1 (string_length "c.bnez"))) with
+ | _s3753_ =>
+ (spc_matches_prefix _s3753_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2291_ _)) =>
- (match (string_drop _s2290_ _s2291_) with
- | _s2292_ =>
- (creg_name_matches_prefix _s2292_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3754_ _) =>
+ (match (string_drop _s3753_ _s3754_) with
+ | _s3755_ =>
+ (creg_name_matches_prefix _s3755_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs, existT _ _s2293_ _)) =>
- (match (string_drop _s2292_ _s2293_) with
- | _s2294_ =>
- (sep_matches_prefix _s2294_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs, existT _ _s3756_ _) =>
+ (match (string_drop _s3755_ _s3756_) with
+ | _s3757_ =>
+ (sep_matches_prefix _s3757_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2295_ _)) =>
- match (string_drop _s2294_ _s2295_) with
- | _s2296_ =>
- match (hex_bits_8_matches_prefix _s2296_) with
- | Some ((imm, existT _ _s2297_ _)) =>
- let p0_ := string_drop _s2296_ _s2297_ in
- if ((generic_eq p0_ "")) then Some ((rs, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3758_ _) =>
+ match (string_drop _s3757_ _s3758_) with
+ | _s3759_ =>
+ match (hex_bits_8_matches_prefix _s3759_) with
+ | Some (imm, existT _ _s3760_ _) =>
+ let p0_ := string_drop _s3759_ _s3760_ in
+ if generic_eq p0_ "" then Some (rs, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- else returnm (None : option ((mword 3 * mword 8))))
+ else returnm None)
: M (option ((mword 3 * mword 8))).
-Definition _s2275_ (_s2276_ : string)
-: M (option ((mword 3 * mword 8))) :=
-
- let _s2277_ := _s2276_ in
- (if ((string_startswith _s2277_ "c.beqz")) then
- (match (string_drop _s2277_ (projT1 (string_length "c.beqz"))) with
- | _s2278_ =>
- (spc_matches_prefix _s2278_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3738_ (_s3739_ : string) : M (option ((mword 3 * mword 8))) :=
+ let _s3740_ := _s3739_ in
+ (if string_startswith _s3740_ "c.beqz" then
+ (match (string_drop _s3740_ (projT1 (string_length "c.beqz"))) with
+ | _s3741_ =>
+ (spc_matches_prefix _s3741_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2279_ _)) =>
- (match (string_drop _s2278_ _s2279_) with
- | _s2280_ =>
- (creg_name_matches_prefix _s2280_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3742_ _) =>
+ (match (string_drop _s3741_ _s3742_) with
+ | _s3743_ =>
+ (creg_name_matches_prefix _s3743_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs, existT _ _s2281_ _)) =>
- (match (string_drop _s2280_ _s2281_) with
- | _s2282_ =>
- (sep_matches_prefix _s2282_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs, existT _ _s3744_ _) =>
+ (match (string_drop _s3743_ _s3744_) with
+ | _s3745_ =>
+ (sep_matches_prefix _s3745_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2283_ _)) =>
- match (string_drop _s2282_ _s2283_) with
- | _s2284_ =>
- match (hex_bits_8_matches_prefix _s2284_) with
- | Some ((imm, existT _ _s2285_ _)) =>
- let p0_ := string_drop _s2284_ _s2285_ in
- if ((generic_eq p0_ "")) then Some ((rs, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3746_ _) =>
+ match (string_drop _s3745_ _s3746_) with
+ | _s3747_ =>
+ match (hex_bits_8_matches_prefix _s3747_) with
+ | Some (imm, existT _ _s3748_ _) =>
+ let p0_ := string_drop _s3747_ _s3748_ in
+ if generic_eq p0_ "" then Some (rs, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- else returnm (None : option ((mword 3 * mword 8))))
+ else returnm None)
: M (option ((mword 3 * mword 8))).
-Definition _s2267_ (_s2268_ : string)
-: M (option (mword 11)) :=
-
- let _s2269_ := _s2268_ in
- (if ((string_startswith _s2269_ "c.j")) then
- (match (string_drop _s2269_ (projT1 (string_length "c.j"))) with
- | _s2270_ =>
- (spc_matches_prefix _s2270_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s2271_ _)) =>
- match (string_drop _s2270_ _s2271_) with
- | _s2272_ =>
- match (hex_bits_11_matches_prefix _s2272_) with
- | Some ((imm, existT _ _s2273_ _)) =>
- let p0_ := string_drop _s2272_ _s2273_ in
- if ((generic_eq p0_ "")) then Some (imm)
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 11))
+Definition _s3730_ (_s3731_ : string) : M (option (mword 11)) :=
+ let _s3732_ := _s3731_ in
+ (if string_startswith _s3732_ "c.j" then
+ (match (string_drop _s3732_ (projT1 (string_length "c.j"))) with
+ | _s3733_ =>
+ (spc_matches_prefix _s3733_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s3734_ _) =>
+ match (string_drop _s3733_ _s3734_) with
+ | _s3735_ =>
+ match (hex_bits_11_matches_prefix _s3735_) with
+ | Some (imm, existT _ _s3736_ _) =>
+ let p0_ := string_drop _s3735_ _s3736_ in
+ if generic_eq p0_ "" then Some imm
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 11))
- else returnm (None : option (mword 11)))
+ else returnm None)
: M (option (mword 11)).
-Definition _s2255_ (_s2256_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s2257_ := _s2256_ in
- (if ((string_startswith _s2257_ "c.addw")) then
- (match (string_drop _s2257_ (projT1 (string_length "c.addw"))) with
- | _s2258_ =>
- (spc_matches_prefix _s2258_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3718_ (_s3719_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s3720_ := _s3719_ in
+ (if string_startswith _s3720_ "c.addw" then
+ (match (string_drop _s3720_ (projT1 (string_length "c.addw"))) with
+ | _s3721_ =>
+ (spc_matches_prefix _s3721_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2259_ _)) =>
- (match (string_drop _s2258_ _s2259_) with
- | _s2260_ =>
- (creg_name_matches_prefix _s2260_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3722_ _) =>
+ (match (string_drop _s3721_ _s3722_) with
+ | _s3723_ =>
+ (creg_name_matches_prefix _s3723_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2261_ _)) =>
- (match (string_drop _s2260_ _s2261_) with
- | _s2262_ =>
- (sep_matches_prefix _s2262_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3724_ _) =>
+ (match (string_drop _s3723_ _s3724_) with
+ | _s3725_ =>
+ (sep_matches_prefix _s3725_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2263_ _)) =>
- (match (string_drop _s2262_ _s2263_) with
- | _s2264_ =>
- (creg_name_matches_prefix _s2264_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3726_ _) =>
+ (match (string_drop _s3725_ _s3726_) with
+ | _s3727_ =>
+ (creg_name_matches_prefix _s3727_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2265_ _)) =>
- let p0_ := string_drop _s2264_ _s2265_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3728_ _) =>
+ let p0_ := string_drop _s3727_ _s3728_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s2243_ (_s2244_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s2245_ := _s2244_ in
- (if ((string_startswith _s2245_ "c.subw")) then
- (match (string_drop _s2245_ (projT1 (string_length "c.subw"))) with
- | _s2246_ =>
- (spc_matches_prefix _s2246_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3706_ (_s3707_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s3708_ := _s3707_ in
+ (if string_startswith _s3708_ "c.subw" then
+ (match (string_drop _s3708_ (projT1 (string_length "c.subw"))) with
+ | _s3709_ =>
+ (spc_matches_prefix _s3709_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2247_ _)) =>
- (match (string_drop _s2246_ _s2247_) with
- | _s2248_ =>
- (creg_name_matches_prefix _s2248_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3710_ _) =>
+ (match (string_drop _s3709_ _s3710_) with
+ | _s3711_ =>
+ (creg_name_matches_prefix _s3711_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2249_ _)) =>
- (match (string_drop _s2248_ _s2249_) with
- | _s2250_ =>
- (sep_matches_prefix _s2250_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3712_ _) =>
+ (match (string_drop _s3711_ _s3712_) with
+ | _s3713_ =>
+ (sep_matches_prefix _s3713_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2251_ _)) =>
- (match (string_drop _s2250_ _s2251_) with
- | _s2252_ =>
- (creg_name_matches_prefix _s2252_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3714_ _) =>
+ (match (string_drop _s3713_ _s3714_) with
+ | _s3715_ =>
+ (creg_name_matches_prefix _s3715_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2253_ _)) =>
- let p0_ := string_drop _s2252_ _s2253_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3716_ _) =>
+ let p0_ := string_drop _s3715_ _s3716_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s2231_ (_s2232_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s2233_ := _s2232_ in
- (if ((string_startswith _s2233_ "c.and")) then
- (match (string_drop _s2233_ (projT1 (string_length "c.and"))) with
- | _s2234_ =>
- (spc_matches_prefix _s2234_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3694_ (_s3695_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s3696_ := _s3695_ in
+ (if string_startswith _s3696_ "c.and" then
+ (match (string_drop _s3696_ (projT1 (string_length "c.and"))) with
+ | _s3697_ =>
+ (spc_matches_prefix _s3697_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2235_ _)) =>
- (match (string_drop _s2234_ _s2235_) with
- | _s2236_ =>
- (creg_name_matches_prefix _s2236_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3698_ _) =>
+ (match (string_drop _s3697_ _s3698_) with
+ | _s3699_ =>
+ (creg_name_matches_prefix _s3699_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2237_ _)) =>
- (match (string_drop _s2236_ _s2237_) with
- | _s2238_ =>
- (sep_matches_prefix _s2238_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3700_ _) =>
+ (match (string_drop _s3699_ _s3700_) with
+ | _s3701_ =>
+ (sep_matches_prefix _s3701_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2239_ _)) =>
- (match (string_drop _s2238_ _s2239_) with
- | _s2240_ =>
- (creg_name_matches_prefix _s2240_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3702_ _) =>
+ (match (string_drop _s3701_ _s3702_) with
+ | _s3703_ =>
+ (creg_name_matches_prefix _s3703_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2241_ _)) =>
- let p0_ := string_drop _s2240_ _s2241_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3704_ _) =>
+ let p0_ := string_drop _s3703_ _s3704_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s2219_ (_s2220_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s2221_ := _s2220_ in
- (if ((string_startswith _s2221_ "c.or")) then
- (match (string_drop _s2221_ (projT1 (string_length "c.or"))) with
- | _s2222_ =>
- (spc_matches_prefix _s2222_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3682_ (_s3683_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s3684_ := _s3683_ in
+ (if string_startswith _s3684_ "c.or" then
+ (match (string_drop _s3684_ (projT1 (string_length "c.or"))) with
+ | _s3685_ =>
+ (spc_matches_prefix _s3685_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2223_ _)) =>
- (match (string_drop _s2222_ _s2223_) with
- | _s2224_ =>
- (creg_name_matches_prefix _s2224_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3686_ _) =>
+ (match (string_drop _s3685_ _s3686_) with
+ | _s3687_ =>
+ (creg_name_matches_prefix _s3687_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2225_ _)) =>
- (match (string_drop _s2224_ _s2225_) with
- | _s2226_ =>
- (sep_matches_prefix _s2226_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3688_ _) =>
+ (match (string_drop _s3687_ _s3688_) with
+ | _s3689_ =>
+ (sep_matches_prefix _s3689_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2227_ _)) =>
- (match (string_drop _s2226_ _s2227_) with
- | _s2228_ =>
- (creg_name_matches_prefix _s2228_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3690_ _) =>
+ (match (string_drop _s3689_ _s3690_) with
+ | _s3691_ =>
+ (creg_name_matches_prefix _s3691_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2229_ _)) =>
- let p0_ := string_drop _s2228_ _s2229_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3692_ _) =>
+ let p0_ := string_drop _s3691_ _s3692_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s2207_ (_s2208_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s2209_ := _s2208_ in
- (if ((string_startswith _s2209_ "c.xor")) then
- (match (string_drop _s2209_ (projT1 (string_length "c.xor"))) with
- | _s2210_ =>
- (spc_matches_prefix _s2210_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3670_ (_s3671_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s3672_ := _s3671_ in
+ (if string_startswith _s3672_ "c.xor" then
+ (match (string_drop _s3672_ (projT1 (string_length "c.xor"))) with
+ | _s3673_ =>
+ (spc_matches_prefix _s3673_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2211_ _)) =>
- (match (string_drop _s2210_ _s2211_) with
- | _s2212_ =>
- (creg_name_matches_prefix _s2212_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3674_ _) =>
+ (match (string_drop _s3673_ _s3674_) with
+ | _s3675_ =>
+ (creg_name_matches_prefix _s3675_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2213_ _)) =>
- (match (string_drop _s2212_ _s2213_) with
- | _s2214_ =>
- (sep_matches_prefix _s2214_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3676_ _) =>
+ (match (string_drop _s3675_ _s3676_) with
+ | _s3677_ =>
+ (sep_matches_prefix _s3677_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2215_ _)) =>
- (match (string_drop _s2214_ _s2215_) with
- | _s2216_ =>
- (creg_name_matches_prefix _s2216_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3678_ _) =>
+ (match (string_drop _s3677_ _s3678_) with
+ | _s3679_ =>
+ (creg_name_matches_prefix _s3679_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2217_ _)) =>
- let p0_ := string_drop _s2216_ _s2217_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3680_ _) =>
+ let p0_ := string_drop _s3679_ _s3680_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s2195_ (_s2196_ : string)
-: M (option ((mword 3 * mword 3))) :=
-
- let _s2197_ := _s2196_ in
- (if ((string_startswith _s2197_ "c.sub")) then
- (match (string_drop _s2197_ (projT1 (string_length "c.sub"))) with
- | _s2198_ =>
- (spc_matches_prefix _s2198_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3658_ (_s3659_ : string) : M (option ((mword 3 * mword 3))) :=
+ let _s3660_ := _s3659_ in
+ (if string_startswith _s3660_ "c.sub" then
+ (match (string_drop _s3660_ (projT1 (string_length "c.sub"))) with
+ | _s3661_ =>
+ (spc_matches_prefix _s3661_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2199_ _)) =>
- (match (string_drop _s2198_ _s2199_) with
- | _s2200_ =>
- (creg_name_matches_prefix _s2200_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3662_ _) =>
+ (match (string_drop _s3661_ _s3662_) with
+ | _s3663_ =>
+ (creg_name_matches_prefix _s3663_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2201_ _)) =>
- (match (string_drop _s2200_ _s2201_) with
- | _s2202_ =>
- (sep_matches_prefix _s2202_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3664_ _) =>
+ (match (string_drop _s3663_ _s3664_) with
+ | _s3665_ =>
+ (sep_matches_prefix _s3665_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2203_ _)) =>
- (match (string_drop _s2202_ _s2203_) with
- | _s2204_ =>
- (creg_name_matches_prefix _s2204_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3666_ _) =>
+ (match (string_drop _s3665_ _s3666_) with
+ | _s3667_ =>
+ (creg_name_matches_prefix _s3667_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2205_ _)) =>
- let p0_ := string_drop _s2204_ _s2205_ in
- if ((generic_eq p0_ "")) then Some ((rsd, rs2))
- else None
- | _ => None
- end)
- : option ((mword 3 * mword 3)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3668_ _) =>
+ let p0_ := string_drop _s3667_ _s3668_ in
+ if generic_eq p0_ "" then Some (rsd, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- | _ => returnm (None : option ((mword 3 * mword 3)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3)))
end)
: M (option ((mword 3 * mword 3)))
- else returnm (None : option ((mword 3 * mword 3))))
+ else returnm None)
: M (option ((mword 3 * mword 3))).
-Definition _s2183_ (_s2184_ : string)
-: M (option ((mword 3 * mword 6))) :=
-
- let _s2185_ := _s2184_ in
- (if ((string_startswith _s2185_ "c.andi")) then
- (match (string_drop _s2185_ (projT1 (string_length "c.andi"))) with
- | _s2186_ =>
- (spc_matches_prefix _s2186_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3646_ (_s3647_ : string) : M (option ((mword 3 * mword 6))) :=
+ let _s3648_ := _s3647_ in
+ (if string_startswith _s3648_ "c.andi" then
+ (match (string_drop _s3648_ (projT1 (string_length "c.andi"))) with
+ | _s3649_ =>
+ (spc_matches_prefix _s3649_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2187_ _)) =>
- (match (string_drop _s2186_ _s2187_) with
- | _s2188_ =>
- (creg_name_matches_prefix _s2188_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3650_ _) =>
+ (match (string_drop _s3649_ _s3650_) with
+ | _s3651_ =>
+ (creg_name_matches_prefix _s3651_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2189_ _)) =>
- (match (string_drop _s2188_ _s2189_) with
- | _s2190_ =>
- (sep_matches_prefix _s2190_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3652_ _) =>
+ (match (string_drop _s3651_ _s3652_) with
+ | _s3653_ =>
+ (sep_matches_prefix _s3653_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2191_ _)) =>
- match (string_drop _s2190_ _s2191_) with
- | _s2192_ =>
- match (hex_bits_6_matches_prefix _s2192_) with
- | Some ((imm, existT _ _s2193_ _)) =>
- let p0_ := string_drop _s2192_ _s2193_ in
- if ((generic_eq p0_ "")) then Some ((rsd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3654_ _) =>
+ match (string_drop _s3653_ _s3654_) with
+ | _s3655_ =>
+ match (hex_bits_6_matches_prefix _s3655_) with
+ | Some (imm, existT _ _s3656_ _) =>
+ let p0_ := string_drop _s3655_ _s3656_ in
+ if generic_eq p0_ "" then Some (rsd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- else returnm (None : option ((mword 3 * mword 6))))
+ else returnm None)
: M (option ((mword 3 * mword 6))).
-Definition _s2171_ (_s2172_ : string)
-: M (option ((mword 3 * mword 6))) :=
-
- let _s2173_ := _s2172_ in
- (if ((string_startswith _s2173_ "c.srai")) then
- (match (string_drop _s2173_ (projT1 (string_length "c.srai"))) with
- | _s2174_ =>
- (spc_matches_prefix _s2174_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3634_ (_s3635_ : string) : M (option ((mword 3 * mword 6))) :=
+ let _s3636_ := _s3635_ in
+ (if string_startswith _s3636_ "c.srai" then
+ (match (string_drop _s3636_ (projT1 (string_length "c.srai"))) with
+ | _s3637_ =>
+ (spc_matches_prefix _s3637_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2175_ _)) =>
- (match (string_drop _s2174_ _s2175_) with
- | _s2176_ =>
- (creg_name_matches_prefix _s2176_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3638_ _) =>
+ (match (string_drop _s3637_ _s3638_) with
+ | _s3639_ =>
+ (creg_name_matches_prefix _s3639_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2177_ _)) =>
- (match (string_drop _s2176_ _s2177_) with
- | _s2178_ =>
- (sep_matches_prefix _s2178_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3640_ _) =>
+ (match (string_drop _s3639_ _s3640_) with
+ | _s3641_ =>
+ (sep_matches_prefix _s3641_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2179_ _)) =>
- match (string_drop _s2178_ _s2179_) with
- | _s2180_ =>
- match (hex_bits_6_matches_prefix _s2180_) with
- | Some ((shamt, existT _ _s2181_ _)) =>
- let p0_ := string_drop _s2180_ _s2181_ in
- if ((generic_eq p0_ "")) then Some ((rsd, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3642_ _) =>
+ match (string_drop _s3641_ _s3642_) with
+ | _s3643_ =>
+ match (hex_bits_6_matches_prefix _s3643_) with
+ | Some (shamt, existT _ _s3644_ _) =>
+ let p0_ := string_drop _s3643_ _s3644_ in
+ if generic_eq p0_ "" then Some (rsd, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- else returnm (None : option ((mword 3 * mword 6))))
+ else returnm None)
: M (option ((mword 3 * mword 6))).
-Definition _s2159_ (_s2160_ : string)
-: M (option ((mword 3 * mword 6))) :=
-
- let _s2161_ := _s2160_ in
- (if ((string_startswith _s2161_ "c.srli")) then
- (match (string_drop _s2161_ (projT1 (string_length "c.srli"))) with
- | _s2162_ =>
- (spc_matches_prefix _s2162_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3622_ (_s3623_ : string) : M (option ((mword 3 * mword 6))) :=
+ let _s3624_ := _s3623_ in
+ (if string_startswith _s3624_ "c.srli" then
+ (match (string_drop _s3624_ (projT1 (string_length "c.srli"))) with
+ | _s3625_ =>
+ (spc_matches_prefix _s3625_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2163_ _)) =>
- (match (string_drop _s2162_ _s2163_) with
- | _s2164_ =>
- (creg_name_matches_prefix _s2164_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3626_ _) =>
+ (match (string_drop _s3625_ _s3626_) with
+ | _s3627_ =>
+ (creg_name_matches_prefix _s3627_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2165_ _)) =>
- (match (string_drop _s2164_ _s2165_) with
- | _s2166_ =>
- (sep_matches_prefix _s2166_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3628_ _) =>
+ (match (string_drop _s3627_ _s3628_) with
+ | _s3629_ =>
+ (sep_matches_prefix _s3629_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2167_ _)) =>
- match (string_drop _s2166_ _s2167_) with
- | _s2168_ =>
- match (hex_bits_6_matches_prefix _s2168_) with
- | Some ((shamt, existT _ _s2169_ _)) =>
- let p0_ := string_drop _s2168_ _s2169_ in
- if ((generic_eq p0_ "")) then Some ((rsd, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3630_ _) =>
+ match (string_drop _s3629_ _s3630_) with
+ | _s3631_ =>
+ match (hex_bits_6_matches_prefix _s3631_) with
+ | Some (shamt, existT _ _s3632_ _) =>
+ let p0_ := string_drop _s3631_ _s3632_ in
+ if generic_eq p0_ "" then Some (rsd, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- | _ => returnm (None : option ((mword 3 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6)))
end)
: M (option ((mword 3 * mword 6)))
- else returnm (None : option ((mword 3 * mword 6))))
+ else returnm None)
: M (option ((mword 3 * mword 6))).
-Definition _s2147_ (_s2148_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2149_ := _s2148_ in
- (if ((string_startswith _s2149_ "c.lui")) then
- (match (string_drop _s2149_ (projT1 (string_length "c.lui"))) with
- | _s2150_ =>
- (spc_matches_prefix _s2150_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3610_ (_s3611_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3612_ := _s3611_ in
+ (if string_startswith _s3612_ "c.lui" then
+ (match (string_drop _s3612_ (projT1 (string_length "c.lui"))) with
+ | _s3613_ =>
+ (spc_matches_prefix _s3613_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2151_ _)) =>
- (match (string_drop _s2150_ _s2151_) with
- | _s2152_ =>
- (reg_name_matches_prefix _s2152_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3614_ _) =>
+ (match (string_drop _s3613_ _s3614_) with
+ | _s3615_ =>
+ (reg_name_matches_prefix _s3615_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2153_ _)) =>
- (match (string_drop _s2152_ _s2153_) with
- | _s2154_ =>
- (sep_matches_prefix _s2154_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3616_ _) =>
+ (match (string_drop _s3615_ _s3616_) with
+ | _s3617_ =>
+ (sep_matches_prefix _s3617_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2155_ _)) =>
- match (string_drop _s2154_ _s2155_) with
- | _s2156_ =>
- match (hex_bits_6_matches_prefix _s2156_) with
- | Some ((imm, existT _ _s2157_ _)) =>
- let p0_ := string_drop _s2156_ _s2157_ in
- if ((generic_eq p0_ "")) then Some ((rd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3618_ _) =>
+ match (string_drop _s3617_ _s3618_) with
+ | _s3619_ =>
+ match (hex_bits_6_matches_prefix _s3619_) with
+ | Some (imm, existT _ _s3620_ _) =>
+ let p0_ := string_drop _s3619_ _s3620_ in
+ if generic_eq p0_ "" then Some (rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2139_ (_s2140_ : string)
-: M (option (mword 6)) :=
-
- let _s2141_ := _s2140_ in
- (if ((string_startswith _s2141_ "c.addi16sp")) then
- (match (string_drop _s2141_ (projT1 (string_length "c.addi16sp"))) with
- | _s2142_ =>
- (spc_matches_prefix _s2142_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s2143_ _)) =>
- match (string_drop _s2142_ _s2143_) with
- | _s2144_ =>
- match (hex_bits_6_matches_prefix _s2144_) with
- | Some ((imm, existT _ _s2145_ _)) =>
- let p0_ := string_drop _s2144_ _s2145_ in
- if ((generic_eq p0_ "")) then Some (imm)
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 6))
+Definition _s3602_ (_s3603_ : string) : M (option (mword 6)) :=
+ let _s3604_ := _s3603_ in
+ (if string_startswith _s3604_ "c.addi16sp" then
+ (match (string_drop _s3604_ (projT1 (string_length "c.addi16sp"))) with
+ | _s3605_ =>
+ (spc_matches_prefix _s3605_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s3606_ _) =>
+ match (string_drop _s3605_ _s3606_) with
+ | _s3607_ =>
+ match (hex_bits_6_matches_prefix _s3607_) with
+ | Some (imm, existT _ _s3608_ _) =>
+ let p0_ := string_drop _s3607_ _s3608_ in
+ if generic_eq p0_ "" then Some imm
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 6))
- else returnm (None : option (mword 6)))
+ else returnm None)
: M (option (mword 6)).
-Definition _s2127_ (_s2128_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2129_ := _s2128_ in
- (if ((string_startswith _s2129_ "c.li")) then
- (match (string_drop _s2129_ (projT1 (string_length "c.li"))) with
- | _s2130_ =>
- (spc_matches_prefix _s2130_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3590_ (_s3591_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3592_ := _s3591_ in
+ (if string_startswith _s3592_ "c.li" then
+ (match (string_drop _s3592_ (projT1 (string_length "c.li"))) with
+ | _s3593_ =>
+ (spc_matches_prefix _s3593_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2131_ _)) =>
- (match (string_drop _s2130_ _s2131_) with
- | _s2132_ =>
- (reg_name_matches_prefix _s2132_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3594_ _) =>
+ (match (string_drop _s3593_ _s3594_) with
+ | _s3595_ =>
+ (reg_name_matches_prefix _s3595_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2133_ _)) =>
- (match (string_drop _s2132_ _s2133_) with
- | _s2134_ =>
- (sep_matches_prefix _s2134_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3596_ _) =>
+ (match (string_drop _s3595_ _s3596_) with
+ | _s3597_ =>
+ (sep_matches_prefix _s3597_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2135_ _)) =>
- match (string_drop _s2134_ _s2135_) with
- | _s2136_ =>
- match (hex_bits_6_matches_prefix _s2136_) with
- | Some ((imm, existT _ _s2137_ _)) =>
- let p0_ := string_drop _s2136_ _s2137_ in
- if ((generic_eq p0_ "")) then Some ((rd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3598_ _) =>
+ match (string_drop _s3597_ _s3598_) with
+ | _s3599_ =>
+ match (hex_bits_6_matches_prefix _s3599_) with
+ | Some (imm, existT _ _s3600_ _) =>
+ let p0_ := string_drop _s3599_ _s3600_ in
+ if generic_eq p0_ "" then Some (rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2115_ (_s2116_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2117_ := _s2116_ in
- (if ((string_startswith _s2117_ "c.addiw")) then
- (match (string_drop _s2117_ (projT1 (string_length "c.addiw"))) with
- | _s2118_ =>
- (spc_matches_prefix _s2118_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3578_ (_s3579_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3580_ := _s3579_ in
+ (if string_startswith _s3580_ "c.addiw" then
+ (match (string_drop _s3580_ (projT1 (string_length "c.addiw"))) with
+ | _s3581_ =>
+ (spc_matches_prefix _s3581_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2119_ _)) =>
- (match (string_drop _s2118_ _s2119_) with
- | _s2120_ =>
- (reg_name_matches_prefix _s2120_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3582_ _) =>
+ (match (string_drop _s3581_ _s3582_) with
+ | _s3583_ =>
+ (reg_name_matches_prefix _s3583_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2121_ _)) =>
- (match (string_drop _s2120_ _s2121_) with
- | _s2122_ =>
- (sep_matches_prefix _s2122_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3584_ _) =>
+ (match (string_drop _s3583_ _s3584_) with
+ | _s3585_ =>
+ (sep_matches_prefix _s3585_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2123_ _)) =>
- match (string_drop _s2122_ _s2123_) with
- | _s2124_ =>
- match (hex_bits_6_matches_prefix _s2124_) with
- | Some ((imm, existT _ _s2125_ _)) =>
- let p0_ := string_drop _s2124_ _s2125_ in
- if ((generic_eq p0_ "")) then Some ((rsd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3586_ _) =>
+ match (string_drop _s3585_ _s3586_) with
+ | _s3587_ =>
+ match (hex_bits_6_matches_prefix _s3587_) with
+ | Some (imm, existT _ _s3588_ _) =>
+ let p0_ := string_drop _s3587_ _s3588_ in
+ if generic_eq p0_ "" then Some (rsd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2107_ (_s2108_ : string)
-: M (option (mword 11)) :=
-
- let _s2109_ := _s2108_ in
- (if ((string_startswith _s2109_ "c.jal")) then
- (match (string_drop _s2109_ (projT1 (string_length "c.jal"))) with
- | _s2110_ =>
- (spc_matches_prefix _s2110_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s2111_ _)) =>
- match (string_drop _s2110_ _s2111_) with
- | _s2112_ =>
- match (hex_bits_12_matches_prefix _s2112_) with
- | Some ((v__814, existT _ _s2113_ _)) =>
- if ((eq_vec (subrange_vec_dec v__814 0 0)
- (vec_of_bits [B0] : mword (0 - 0 + 1)))) then
- let imm : mword 11 := subrange_vec_dec v__814 11 1 in
- let imm : mword 11 := subrange_vec_dec v__814 11 1 in
- let p0_ := string_drop _s2112_ _s2113_ in
- if ((generic_eq p0_ "")) then Some (imm)
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option (mword 11))
+Definition _s3570_ (_s3571_ : string) : M (option (mword 11)) :=
+ let _s3572_ := _s3571_ in
+ (if string_startswith _s3572_ "c.jal" then
+ (match (string_drop _s3572_ (projT1 (string_length "c.jal"))) with
+ | _s3573_ =>
+ (spc_matches_prefix _s3573_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s3574_ _) =>
+ match (string_drop _s3573_ _s3574_) with
+ | _s3575_ =>
+ match (hex_bits_12_matches_prefix _s3575_) with
+ | Some (v__1366, existT _ _s3576_ _) =>
+ if eq_vec (subrange_vec_dec v__1366 0 0) ('b"0" : mword (0 - 0 + 1))
+ then
+ let imm : mword 11 := subrange_vec_dec v__1366 11 1 in
+ let imm : mword 11 := subrange_vec_dec v__1366 11 1 in
+ let p0_ := string_drop _s3575_ _s3576_ in
+ if generic_eq p0_ "" then Some imm
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option (mword 11))
- else returnm (None : option (mword 11)))
+ else returnm None)
: M (option (mword 11)).
-Definition _s2095_ (_s2096_ : string)
-: M (option ((mword 5 * mword 6))) :=
-
- let _s2097_ := _s2096_ in
- (if ((string_startswith _s2097_ "c.addi")) then
- (match (string_drop _s2097_ (projT1 (string_length "c.addi"))) with
- | _s2098_ =>
- (spc_matches_prefix _s2098_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3558_ (_s3559_ : string) : M (option ((mword 5 * mword 6))) :=
+ let _s3560_ := _s3559_ in
+ (if string_startswith _s3560_ "c.addi" then
+ (match (string_drop _s3560_ (projT1 (string_length "c.addi"))) with
+ | _s3561_ =>
+ (spc_matches_prefix _s3561_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2099_ _)) =>
- (match (string_drop _s2098_ _s2099_) with
- | _s2100_ =>
- (reg_name_matches_prefix _s2100_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3562_ _) =>
+ (match (string_drop _s3561_ _s3562_) with
+ | _s3563_ =>
+ (reg_name_matches_prefix _s3563_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s2101_ _)) =>
- (match (string_drop _s2100_ _s2101_) with
- | _s2102_ =>
- (sep_matches_prefix _s2102_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s3564_ _) =>
+ (match (string_drop _s3563_ _s3564_) with
+ | _s3565_ =>
+ (sep_matches_prefix _s3565_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2103_ _)) =>
- match (string_drop _s2102_ _s2103_) with
- | _s2104_ =>
- match (hex_bits_6_matches_prefix _s2104_) with
- | Some ((nzi, existT _ _s2105_ _)) =>
- let p0_ := string_drop _s2104_ _s2105_ in
- if ((generic_eq p0_ "")) then Some ((rsd, nzi))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3566_ _) =>
+ match (string_drop _s3565_ _s3566_) with
+ | _s3567_ =>
+ match (hex_bits_6_matches_prefix _s3567_) with
+ | Some (nzi, existT _ _s3568_ _) =>
+ let p0_ := string_drop _s3567_ _s3568_ in
+ if generic_eq p0_ "" then Some (rsd, nzi)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- | _ => returnm (None : option ((mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6)))
end)
: M (option ((mword 5 * mword 6)))
- else returnm (None : option ((mword 5 * mword 6))))
+ else returnm None)
: M (option ((mword 5 * mword 6))).
-Definition _s2079_ (_s2080_ : string)
-: M (option ((mword 3 * mword 3 * mword 5))) :=
-
- let _s2081_ := _s2080_ in
- (if ((string_startswith _s2081_ "c.sd")) then
- (match (string_drop _s2081_ (projT1 (string_length "c.sd"))) with
- | _s2082_ =>
- (spc_matches_prefix _s2082_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3542_ (_s3543_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s3544_ := _s3543_ in
+ (if string_startswith _s3544_ "c.sd" then
+ (match (string_drop _s3544_ (projT1 (string_length "c.sd"))) with
+ | _s3545_ =>
+ (spc_matches_prefix _s3545_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2083_ _)) =>
- (match (string_drop _s2082_ _s2083_) with
- | _s2084_ =>
- (creg_name_matches_prefix _s2084_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3546_ _) =>
+ (match (string_drop _s3545_ _s3546_) with
+ | _s3547_ =>
+ (creg_name_matches_prefix _s3547_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsc1, existT _ _s2085_ _)) =>
- (match (string_drop _s2084_ _s2085_) with
- | _s2086_ =>
- (sep_matches_prefix _s2086_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc1, existT _ _s3548_ _) =>
+ (match (string_drop _s3547_ _s3548_) with
+ | _s3549_ =>
+ (sep_matches_prefix _s3549_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2087_ _)) =>
- (match (string_drop _s2086_ _s2087_) with
- | _s2088_ =>
- (creg_name_matches_prefix _s2088_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3550_ _) =>
+ (match (string_drop _s3549_ _s3550_) with
+ | _s3551_ =>
+ (creg_name_matches_prefix _s3551_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc2, existT _ _s2089_ _)) =>
- (match (string_drop _s2088_ _s2089_) with
- | _s2090_ =>
- (sep_matches_prefix _s2090_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc2, existT _ _s3552_ _) =>
+ (match (string_drop _s3551_ _s3552_) with
+ | _s3553_ =>
+ (sep_matches_prefix _s3553_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2091_ _)) =>
- match (string_drop _s2090_ _s2091_) with
- | _s2092_ =>
- match (hex_bits_8_matches_prefix _s2092_) with
- | Some ((v__816, existT _ _s2093_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__816 2 0)
- (vec_of_bits [B0;B0;B0]
- : mword (2 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__816 7 3 in
- let uimm : mword 5 :=
- subrange_vec_dec v__816 7 3 in
- let p0_ :=
- string_drop _s2092_ _s2093_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rsc1, rsc2, uimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s3554_ _) =>
+ match (string_drop _s3553_ _s3554_) with
+ | _s3555_ =>
+ match (hex_bits_8_matches_prefix _s3555_) with
+ | Some (v__1368, existT _ _s3556_ _) =>
+ if eq_vec (subrange_vec_dec v__1368 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1368 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1368 7 3 in
+ let p0_ := string_drop _s3555_ _s3556_ in
+ if generic_eq p0_ "" then
+ Some (rsc1, rsc2, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5))).
-Definition _s2063_ (_s2064_ : string)
-: M (option ((mword 3 * mword 3 * mword 5))) :=
-
- let _s2065_ := _s2064_ in
- (if ((string_startswith _s2065_ "c.sw")) then
- (match (string_drop _s2065_ (projT1 (string_length "c.sw"))) with
- | _s2066_ =>
- (spc_matches_prefix _s2066_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3526_ (_s3527_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s3528_ := _s3527_ in
+ (if string_startswith _s3528_ "c.sw" then
+ (match (string_drop _s3528_ (projT1 (string_length "c.sw"))) with
+ | _s3529_ =>
+ (spc_matches_prefix _s3529_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2067_ _)) =>
- (match (string_drop _s2066_ _s2067_) with
- | _s2068_ =>
- (creg_name_matches_prefix _s2068_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3530_ _) =>
+ (match (string_drop _s3529_ _s3530_) with
+ | _s3531_ =>
+ (creg_name_matches_prefix _s3531_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsc1, existT _ _s2069_ _)) =>
- (match (string_drop _s2068_ _s2069_) with
- | _s2070_ =>
- (sep_matches_prefix _s2070_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc1, existT _ _s3532_ _) =>
+ (match (string_drop _s3531_ _s3532_) with
+ | _s3533_ =>
+ (sep_matches_prefix _s3533_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2071_ _)) =>
- (match (string_drop _s2070_ _s2071_) with
- | _s2072_ =>
- (creg_name_matches_prefix _s2072_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3534_ _) =>
+ (match (string_drop _s3533_ _s3534_) with
+ | _s3535_ =>
+ (creg_name_matches_prefix _s3535_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc2, existT _ _s2073_ _)) =>
- (match (string_drop _s2072_ _s2073_) with
- | _s2074_ =>
- (sep_matches_prefix _s2074_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc2, existT _ _s3536_ _) =>
+ (match (string_drop _s3535_ _s3536_) with
+ | _s3537_ =>
+ (sep_matches_prefix _s3537_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2075_ _)) =>
- match (string_drop _s2074_ _s2075_) with
- | _s2076_ =>
- match (hex_bits_7_matches_prefix _s2076_) with
- | Some ((v__818, existT _ _s2077_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__818 1 0)
- (vec_of_bits [B0;B0]
- : mword (1 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__818 6 2 in
- let uimm : mword 5 :=
- subrange_vec_dec v__818 6 2 in
- let p0_ :=
- string_drop _s2076_ _s2077_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rsc1, rsc2, uimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s3538_ _) =>
+ match (string_drop _s3537_ _s3538_) with
+ | _s3539_ =>
+ match (hex_bits_7_matches_prefix _s3539_) with
+ | Some (v__1370, existT _ _s3540_ _) =>
+ if eq_vec (subrange_vec_dec v__1370 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1370 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1370 6 2 in
+ let p0_ := string_drop _s3539_ _s3540_ in
+ if generic_eq p0_ "" then
+ Some (rsc1, rsc2, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5))).
-Definition _s2047_ (_s2048_ : string)
-: M (option ((mword 3 * mword 3 * mword 5))) :=
-
- let _s2049_ := _s2048_ in
- (if ((string_startswith _s2049_ "c.ld")) then
- (match (string_drop _s2049_ (projT1 (string_length "c.ld"))) with
- | _s2050_ =>
- (spc_matches_prefix _s2050_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3510_ (_s3511_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s3512_ := _s3511_ in
+ (if string_startswith _s3512_ "c.ld" then
+ (match (string_drop _s3512_ (projT1 (string_length "c.ld"))) with
+ | _s3513_ =>
+ (spc_matches_prefix _s3513_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2051_ _)) =>
- (match (string_drop _s2050_ _s2051_) with
- | _s2052_ =>
- (creg_name_matches_prefix _s2052_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3514_ _) =>
+ (match (string_drop _s3513_ _s3514_) with
+ | _s3515_ =>
+ (creg_name_matches_prefix _s3515_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s2053_ _)) =>
- (match (string_drop _s2052_ _s2053_) with
- | _s2054_ =>
- (sep_matches_prefix _s2054_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s3516_ _) =>
+ (match (string_drop _s3515_ _s3516_) with
+ | _s3517_ =>
+ (sep_matches_prefix _s3517_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2055_ _)) =>
- (match (string_drop _s2054_ _s2055_) with
- | _s2056_ =>
- (creg_name_matches_prefix _s2056_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3518_ _) =>
+ (match (string_drop _s3517_ _s3518_) with
+ | _s3519_ =>
+ (creg_name_matches_prefix _s3519_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc, existT _ _s2057_ _)) =>
- (match (string_drop _s2056_ _s2057_) with
- | _s2058_ =>
- (sep_matches_prefix _s2058_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc, existT _ _s3520_ _) =>
+ (match (string_drop _s3519_ _s3520_) with
+ | _s3521_ =>
+ (sep_matches_prefix _s3521_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2059_ _)) =>
- match (string_drop _s2058_ _s2059_) with
- | _s2060_ =>
- match (hex_bits_8_matches_prefix _s2060_) with
- | Some ((v__820, existT _ _s2061_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__820 2 0)
- (vec_of_bits [B0;B0;B0]
- : mword (2 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__820 7 3 in
- let uimm : mword 5 :=
- subrange_vec_dec v__820 7 3 in
- let p0_ :=
- string_drop _s2060_ _s2061_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rdc, rsc, uimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s3522_ _) =>
+ match (string_drop _s3521_ _s3522_) with
+ | _s3523_ =>
+ match (hex_bits_8_matches_prefix _s3523_) with
+ | Some (v__1372, existT _ _s3524_ _) =>
+ if eq_vec (subrange_vec_dec v__1372 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1372 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1372 7 3 in
+ let p0_ := string_drop _s3523_ _s3524_ in
+ if generic_eq p0_ "" then
+ Some (rdc, rsc, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5))).
-Definition _s2031_ (_s2032_ : string)
-: M (option ((mword 3 * mword 3 * mword 5))) :=
-
- let _s2033_ := _s2032_ in
- (if ((string_startswith _s2033_ "c.lw")) then
- (match (string_drop _s2033_ (projT1 (string_length "c.lw"))) with
- | _s2034_ =>
- (spc_matches_prefix _s2034_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3494_ (_s3495_ : string) : M (option ((mword 3 * mword 3 * mword 5))) :=
+ let _s3496_ := _s3495_ in
+ (if string_startswith _s3496_ "c.lw" then
+ (match (string_drop _s3496_ (projT1 (string_length "c.lw"))) with
+ | _s3497_ =>
+ (spc_matches_prefix _s3497_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2035_ _)) =>
- (match (string_drop _s2034_ _s2035_) with
- | _s2036_ =>
- (creg_name_matches_prefix _s2036_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3498_ _) =>
+ (match (string_drop _s3497_ _s3498_) with
+ | _s3499_ =>
+ (creg_name_matches_prefix _s3499_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s2037_ _)) =>
- (match (string_drop _s2036_ _s2037_) with
- | _s2038_ =>
- (sep_matches_prefix _s2038_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s3500_ _) =>
+ (match (string_drop _s3499_ _s3500_) with
+ | _s3501_ =>
+ (sep_matches_prefix _s3501_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2039_ _)) =>
- (match (string_drop _s2038_ _s2039_) with
- | _s2040_ =>
- (creg_name_matches_prefix _s2040_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3502_ _) =>
+ (match (string_drop _s3501_ _s3502_) with
+ | _s3503_ =>
+ (creg_name_matches_prefix _s3503_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc, existT _ _s2041_ _)) =>
- (match (string_drop _s2040_ _s2041_) with
- | _s2042_ =>
- (sep_matches_prefix _s2042_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc, existT _ _s3504_ _) =>
+ (match (string_drop _s3503_ _s3504_) with
+ | _s3505_ =>
+ (sep_matches_prefix _s3505_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2043_ _)) =>
- match (string_drop _s2042_ _s2043_) with
- | _s2044_ =>
- match (hex_bits_7_matches_prefix _s2044_) with
- | Some ((v__822, existT _ _s2045_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__822 1 0)
- (vec_of_bits [B0;B0]
- : mword (1 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__822 6 2 in
- let uimm : mword 5 :=
- subrange_vec_dec v__822 6 2 in
- let p0_ :=
- string_drop _s2044_ _s2045_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rdc, rsc, uimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s3506_ _) =>
+ match (string_drop _s3505_ _s3506_) with
+ | _s3507_ =>
+ match (hex_bits_7_matches_prefix _s3507_) with
+ | Some (v__1374, existT _ _s3508_ _) =>
+ if eq_vec (subrange_vec_dec v__1374 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1374 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1374 6 2 in
+ let p0_ := string_drop _s3507_ _s3508_ in
+ if generic_eq p0_ "" then
+ Some (rdc, rsc, uimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5)))
end)
: M (option ((mword 3 * mword 3 * mword 5)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5))).
-Definition _s2019_ (_s2020_ : string)
-: M (option ((mword 3 * mword 8))) :=
-
- let _s2021_ := _s2020_ in
- (if ((string_startswith _s2021_ "c.addi4spn")) then
- (match (string_drop _s2021_ (projT1 (string_length "c.addi4spn"))) with
- | _s2022_ =>
- (spc_matches_prefix _s2022_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3482_ (_s3483_ : string) : M (option ((mword 3 * mword 8))) :=
+ let _s3484_ := _s3483_ in
+ (if string_startswith _s3484_ "c.addi4spn" then
+ (match (string_drop _s3484_ (projT1 (string_length "c.addi4spn"))) with
+ | _s3485_ =>
+ (spc_matches_prefix _s3485_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2023_ _)) =>
- (match (string_drop _s2022_ _s2023_) with
- | _s2024_ =>
- (creg_name_matches_prefix _s2024_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3486_ _) =>
+ (match (string_drop _s3485_ _s3486_) with
+ | _s3487_ =>
+ (creg_name_matches_prefix _s3487_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s2025_ _)) =>
- (match (string_drop _s2024_ _s2025_) with
- | _s2026_ =>
- (sep_matches_prefix _s2026_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s3488_ _) =>
+ (match (string_drop _s3487_ _s3488_) with
+ | _s3489_ =>
+ (sep_matches_prefix _s3489_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2027_ _)) =>
- match (string_drop _s2026_ _s2027_) with
- | _s2028_ =>
- match (hex_bits_10_matches_prefix _s2028_) with
- | Some ((v__824, existT _ _s2029_ _)) =>
- if ((eq_vec (subrange_vec_dec v__824 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1))))
- then
- let nzimm : mword 8 := subrange_vec_dec v__824 9 2 in
- let nzimm : mword 8 := subrange_vec_dec v__824 9 2 in
- let p0_ := string_drop _s2028_ _s2029_ in
- if ((generic_eq p0_ "")) then Some ((rdc, nzimm))
- else None
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3490_ _) =>
+ match (string_drop _s3489_ _s3490_) with
+ | _s3491_ =>
+ match (hex_bits_10_matches_prefix _s3491_) with
+ | Some (v__1376, existT _ _s3492_ _) =>
+ if eq_vec (subrange_vec_dec v__1376 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let nzimm : mword 8 := subrange_vec_dec v__1376 9 2 in
+ let nzimm : mword 8 := subrange_vec_dec v__1376 9 2 in
+ let p0_ := string_drop _s3491_ _s3492_ in
+ if generic_eq p0_ "" then Some (rdc, nzimm)
+ else None
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- | _ => returnm (None : option ((mword 3 * mword 8)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8)))
end)
: M (option ((mword 3 * mword 8)))
- else returnm (None : option ((mword 3 * mword 8))))
+ else returnm None)
: M (option ((mword 3 * mword 8))).
-Definition _s1995_ (_s1996_ : string)
+Definition _s3456_ (_s3457_ : string)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1996_ with
- | _s1997_ =>
- (amo_mnemonic_matches_prefix _s1997_) >>= fun w__0 : option ((amoop * {n : Z & ArithFact (n >=
+ (match _s3457_ with
+ | _s3458_ =>
+ (amo_mnemonic_matches_prefix _s3458_) >>= fun w__0 : option ((amoop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1998_ _)) =>
- let _s1999_ := string_drop _s1997_ _s1998_ in
- (if ((string_startswith _s1999_ ".")) then
- (match (string_drop _s1999_ (projT1 (string_length "."))) with
- | _s2000_ =>
- (size_mnemonic_matches_prefix _s2000_) >>= fun w__1 : option ((word_width * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3459_ _) =>
+ let _s3460_ := string_drop _s3458_ _s3459_ in
+ (if string_startswith _s3460_ "." then
+ (match (string_drop _s3460_ (projT1 (string_length "."))) with
+ | _s3461_ =>
+ (size_mnemonic_matches_prefix _s3461_) >>= fun w__1 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((width, existT _ _s2001_ _)) =>
- (match (string_drop _s2000_ _s2001_) with
- | _s2002_ =>
- (maybe_aq_matches_prefix _s2002_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (width, existT _ _s3462_ _) =>
+ (match (string_drop _s3461_ _s3462_) with
+ | _s3463_ =>
+ (maybe_aq_matches_prefix _s3463_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((aq, existT _ _s2003_ _)) =>
- (match (string_drop _s2002_ _s2003_) with
- | _s2004_ =>
- (maybe_rl_matches_prefix _s2004_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s3464_ _) =>
+ (match (string_drop _s3463_ _s3464_) with
+ | _s3465_ =>
+ (maybe_rl_matches_prefix _s3465_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rl, existT _ _s2005_ _)) =>
- (match (string_drop _s2004_ _s2005_) with
- | _s2006_ =>
- (spc_matches_prefix _s2006_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s3466_ _) =>
+ (match (string_drop _s3465_ _s3466_) with
+ | _s3467_ =>
+ (spc_matches_prefix _s3467_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s2007_ _)) =>
- (match (string_drop _s2006_ _s2007_) with
- | _s2008_ =>
- (reg_name_matches_prefix _s2008_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3468_ _) =>
+ (match (string_drop _s3467_ _s3468_) with
+ | _s3469_ =>
+ (reg_name_matches_prefix _s3469_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((rd, existT _ _s2009_ _)) =>
- (match (string_drop _s2008_ _s2009_) with
- | _s2010_ =>
- (sep_matches_prefix _s2010_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3470_ _) =>
+ (match (string_drop _s3469_ _s3470_) with
+ | _s3471_ =>
+ (sep_matches_prefix _s3471_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s2011_ _)) =>
- (match (string_drop _s2010_ _s2011_) with
- | _s2012_ =>
- (reg_name_matches_prefix _s2012_) >>= fun w__7 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3472_ _) =>
+ (match (string_drop _s3471_ _s3472_) with
+ | _s3473_ =>
+ (reg_name_matches_prefix _s3473_) >>= fun w__7 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
- | Some
- ((rs1, existT _ _s2013_ _)) =>
- (match (string_drop _s2012_
- _s2013_) with
- | _s2014_ =>
+ | Some (rs2, existT _ _s3474_ _) =>
+ (match (string_drop _s3473_
+ _s3474_) with
+ | _s3475_ =>
(sep_matches_prefix
- _s2014_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=
+ _s3475_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((tt, existT _ _s2015_ _)) =>
- (match (string_drop
- _s2014_
- _s2015_) with
- | _s2016_ =>
- (reg_name_matches_prefix
- _s2016_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=
- 0)})) =>
- returnm ((match w__9 with
- | Some
- ((rs2, existT _ _s2017_ _)) =>
- let p0_ :=
- string_drop
- _s2016_
- _s2017_ in
- if ((generic_eq
- p0_
- ""))
- then
- Some
- ((op, width, aq, rl, rd, rs1, rs2))
- else
+ (tt, existT _ _s3476_ _) =>
+ let _s3477_ :=
+ string_drop _s3475_
+ _s3476_ in
+ (if string_startswith
+ _s3477_ "(" then
+ (match (string_drop
+ _s3477_
+ (projT1
+ (string_length
+ "("))) with
+ | _s3478_ =>
+ (reg_name_matches_prefix
+ _s3478_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__9 with
+ | Some
+ (rs1, existT _ _s3479_ _) =>
+ let _s3480_ :=
+ string_drop
+ _s3478_
+ _s3479_ in
+ if string_startswith
+ _s3480_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s3480_
+ (projT1
+ (string_length
+ ")")) in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (op, width, aq, rl, rd, rs2, rs1)
+ else
+ None
+ else
+ None
+ | _ =>
None
- | _ =>
- None
- end)
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- end)
+ end)
+ end)
+ : M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ else returnm None)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- else
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5))).
-Definition _s1973_ (_s1974_ : string)
+Definition _s3434_ (_s3435_ : string)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5))) :=
-
- let _s1975_ := _s1974_ in
- (if ((string_startswith _s1975_ "sc.")) then
- (match (string_drop _s1975_ (projT1 (string_length "sc."))) with
- | _s1976_ =>
- (size_mnemonic_matches_prefix _s1976_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s3436_ := _s3435_ in
+ (if string_startswith _s3436_ "sc." then
+ (match (string_drop _s3436_ (projT1 (string_length "sc."))) with
+ | _s3437_ =>
+ (size_mnemonic_matches_prefix _s3437_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s1977_ _)) =>
- (match (string_drop _s1976_ _s1977_) with
- | _s1978_ =>
- (maybe_aq_matches_prefix _s1978_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s3438_ _) =>
+ (match (string_drop _s3437_ _s3438_) with
+ | _s3439_ =>
+ (maybe_aq_matches_prefix _s3439_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s1979_ _)) =>
- (match (string_drop _s1978_ _s1979_) with
- | _s1980_ =>
- (maybe_rl_matches_prefix _s1980_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s3440_ _) =>
+ (match (string_drop _s3439_ _s3440_) with
+ | _s3441_ =>
+ (maybe_rl_matches_prefix _s3441_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s1981_ _)) =>
- (match (string_drop _s1980_ _s1981_) with
- | _s1982_ =>
- (spc_matches_prefix _s1982_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s3442_ _) =>
+ (match (string_drop _s3441_ _s3442_) with
+ | _s3443_ =>
+ (spc_matches_prefix _s3443_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1983_ _)) =>
- (match (string_drop _s1982_ _s1983_) with
- | _s1984_ =>
- (reg_name_matches_prefix _s1984_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3444_ _) =>
+ (match (string_drop _s3443_ _s3444_) with
+ | _s3445_ =>
+ (reg_name_matches_prefix _s3445_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rd, existT _ _s1985_ _)) =>
- (match (string_drop _s1984_ _s1985_) with
- | _s1986_ =>
- (sep_matches_prefix _s1986_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3446_ _) =>
+ (match (string_drop _s3445_ _s3446_) with
+ | _s3447_ =>
+ (sep_matches_prefix _s3447_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1987_ _)) =>
- (match (string_drop _s1986_ _s1987_) with
- | _s1988_ =>
- (reg_name_matches_prefix _s1988_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3448_ _) =>
+ (match (string_drop _s3447_ _s3448_) with
+ | _s3449_ =>
+ (reg_name_matches_prefix _s3449_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((rs1, existT _ _s1989_ _)) =>
- (match (string_drop _s1988_ _s1989_) with
- | _s1990_ =>
- (sep_matches_prefix _s1990_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3450_ _) =>
+ (match (string_drop _s3449_ _s3450_) with
+ | _s3451_ =>
+ (sep_matches_prefix _s3451_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
- | Some ((tt, existT _ _s1991_ _)) =>
- (match (string_drop _s1990_
- _s1991_) with
- | _s1992_ =>
+ | Some (tt, existT _ _s3452_ _) =>
+ (match (string_drop _s3451_
+ _s3452_) with
+ | _s3453_ =>
(reg_name_matches_prefix
- _s1992_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s3453_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__8 with
- | Some
- ((rs2, existT _ _s1993_ _)) =>
- let p0_ :=
- string_drop
- _s1992_
- _s1993_ in
- if ((generic_eq
- p0_ ""))
- then
- Some
- ((size, aq, rl, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ returnm (match w__8 with
+ | Some
+ (rs2, existT _ _s3454_ _) =>
+ let p0_ :=
+ string_drop
+ _s3453_
+ _s3454_ in
+ if generic_eq
+ p0_ ""
+ then
+ Some
+ (size, aq, rl, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)))
- else returnm (None : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5))).
-Definition _s1955_ (_s1956_ : string)
-: M (option ((word_width * bool * bool * mword 5 * mword 5))) :=
-
- let _s1957_ := _s1956_ in
- (if ((string_startswith _s1957_ "lr.")) then
- (match (string_drop _s1957_ (projT1 (string_length "lr."))) with
- | _s1958_ =>
- (size_mnemonic_matches_prefix _s1958_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+Definition _s3416_ (_s3417_ : string) : M (option ((word_width * bool * bool * mword 5 * mword 5))) :=
+ let _s3418_ := _s3417_ in
+ (if string_startswith _s3418_ "lr." then
+ (match (string_drop _s3418_ (projT1 (string_length "lr."))) with
+ | _s3419_ =>
+ (size_mnemonic_matches_prefix _s3419_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s1959_ _)) =>
- (match (string_drop _s1958_ _s1959_) with
- | _s1960_ =>
- (maybe_aq_matches_prefix _s1960_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s3420_ _) =>
+ (match (string_drop _s3419_ _s3420_) with
+ | _s3421_ =>
+ (maybe_aq_matches_prefix _s3421_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s1961_ _)) =>
- (match (string_drop _s1960_ _s1961_) with
- | _s1962_ =>
- (maybe_rl_matches_prefix _s1962_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s3422_ _) =>
+ (match (string_drop _s3421_ _s3422_) with
+ | _s3423_ =>
+ (maybe_rl_matches_prefix _s3423_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s1963_ _)) =>
- (match (string_drop _s1962_ _s1963_) with
- | _s1964_ =>
- (spc_matches_prefix _s1964_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s3424_ _) =>
+ (match (string_drop _s3423_ _s3424_) with
+ | _s3425_ =>
+ (spc_matches_prefix _s3425_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1965_ _)) =>
- (match (string_drop _s1964_ _s1965_) with
- | _s1966_ =>
- (reg_name_matches_prefix _s1966_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3426_ _) =>
+ (match (string_drop _s3425_ _s3426_) with
+ | _s3427_ =>
+ (reg_name_matches_prefix _s3427_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rd, existT _ _s1967_ _)) =>
- (match (string_drop _s1966_ _s1967_) with
- | _s1968_ =>
- (sep_matches_prefix _s1968_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3428_ _) =>
+ (match (string_drop _s3427_ _s3428_) with
+ | _s3429_ =>
+ (sep_matches_prefix _s3429_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1969_ _)) =>
- (match (string_drop _s1968_ _s1969_) with
- | _s1970_ =>
- (reg_name_matches_prefix _s1970_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3430_ _) =>
+ (match (string_drop _s3429_ _s3430_) with
+ | _s3431_ =>
+ (reg_name_matches_prefix _s3431_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs1, existT _ _s1971_ _)) =>
- let p0_ :=
- string_drop _s1970_ _s1971_ in
- if ((generic_eq p0_ "")) then
- Some
- ((size, aq, rl, rd, rs1))
- else None
- | _ => None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs1, existT _ _s3432_ _) =>
+ let p0_ :=
+ string_drop _s3431_ _s3432_ in
+ if generic_eq p0_ "" then
+ Some (size, aq, rl, rd, rs1)
+ else None
+ | _ => None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ => returnm (None : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- | _ => returnm (None : option ((word_width * bool * bool * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5)))
- else returnm (None : option ((word_width * bool * bool * mword 5 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 5))).
-Definition _s1943_ (_s1944_ : string)
-: M (option ((mword 5 * mword 5))) :=
-
- let _s1945_ := _s1944_ in
- (if ((string_startswith _s1945_ "sfence.vma")) then
- (match (string_drop _s1945_ (projT1 (string_length "sfence.vma"))) with
- | _s1946_ =>
- (spc_matches_prefix _s1946_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3404_ (_s3405_ : string) : M (option ((mword 5 * mword 5))) :=
+ let _s3406_ := _s3405_ in
+ (if string_startswith _s3406_ "sfence.vma" then
+ (match (string_drop _s3406_ (projT1 (string_length "sfence.vma"))) with
+ | _s3407_ =>
+ (spc_matches_prefix _s3407_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1947_ _)) =>
- (match (string_drop _s1946_ _s1947_) with
- | _s1948_ =>
- (reg_name_matches_prefix _s1948_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3408_ _) =>
+ (match (string_drop _s3407_ _s3408_) with
+ | _s3409_ =>
+ (reg_name_matches_prefix _s3409_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs1, existT _ _s1949_ _)) =>
- (match (string_drop _s1948_ _s1949_) with
- | _s1950_ =>
- (sep_matches_prefix _s1950_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3410_ _) =>
+ (match (string_drop _s3409_ _s3410_) with
+ | _s3411_ =>
+ (sep_matches_prefix _s3411_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1951_ _)) =>
- (match (string_drop _s1950_ _s1951_) with
- | _s1952_ =>
- (reg_name_matches_prefix _s1952_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3412_ _) =>
+ (match (string_drop _s3411_ _s3412_) with
+ | _s3413_ =>
+ (reg_name_matches_prefix _s3413_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s1953_ _)) =>
- let p0_ := string_drop _s1952_ _s1953_ in
- if ((generic_eq p0_ "")) then Some ((rs1, rs2))
- else None
- | _ => None
- end)
- : option ((mword 5 * mword 5)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s3414_ _) =>
+ let p0_ := string_drop _s3413_ _s3414_ in
+ if generic_eq p0_ "" then Some (rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- | _ => returnm (None : option ((mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5)))
end)
: M (option ((mword 5 * mword 5)))
- else returnm (None : option ((mword 5 * mword 5))))
+ else returnm None)
: M (option ((mword 5 * mword 5))).
-Definition _s1931_ (_s1932_ : string)
-: M (option ((mword 4 * mword 4))) :=
-
- let _s1933_ := _s1932_ in
- (if ((string_startswith _s1933_ "fence.tso")) then
- (match (string_drop _s1933_ (projT1 (string_length "fence.tso"))) with
- | _s1934_ =>
- (spc_matches_prefix _s1934_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3392_ (_s3393_ : string) : M (option ((mword 4 * mword 4))) :=
+ let _s3394_ := _s3393_ in
+ (if string_startswith _s3394_ "fence.tso" then
+ (match (string_drop _s3394_ (projT1 (string_length "fence.tso"))) with
+ | _s3395_ =>
+ (spc_matches_prefix _s3395_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1935_ _)) =>
- (match (string_drop _s1934_ _s1935_) with
- | _s1936_ =>
- (fence_bits_matches_prefix _s1936_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3396_ _) =>
+ (match (string_drop _s3395_ _s3396_) with
+ | _s3397_ =>
+ (fence_bits_matches_prefix _s3397_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((pred, existT _ _s1937_ _)) =>
- (match (string_drop _s1936_ _s1937_) with
- | _s1938_ =>
- (sep_matches_prefix _s1938_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (pred, existT _ _s3398_ _) =>
+ (match (string_drop _s3397_ _s3398_) with
+ | _s3399_ =>
+ (sep_matches_prefix _s3399_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1939_ _)) =>
- (match (string_drop _s1938_ _s1939_) with
- | _s1940_ =>
- (fence_bits_matches_prefix _s1940_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3400_ _) =>
+ (match (string_drop _s3399_ _s3400_) with
+ | _s3401_ =>
+ (fence_bits_matches_prefix _s3401_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((succ, existT _ _s1941_ _)) =>
- let p0_ := string_drop _s1940_ _s1941_ in
- if ((generic_eq p0_ "")) then Some ((pred, succ))
- else None
- | _ => None
- end)
- : option ((mword 4 * mword 4)))
+ returnm (match w__3 with
+ | Some (succ, existT _ _s3402_ _) =>
+ let p0_ := string_drop _s3401_ _s3402_ in
+ if generic_eq p0_ "" then Some (pred, succ)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- else returnm (None : option ((mword 4 * mword 4))))
+ else returnm None)
: M (option ((mword 4 * mword 4))).
-Definition _s1919_ (_s1920_ : string)
-: M (option ((mword 4 * mword 4))) :=
-
- let _s1921_ := _s1920_ in
- (if ((string_startswith _s1921_ "fence")) then
- (match (string_drop _s1921_ (projT1 (string_length "fence"))) with
- | _s1922_ =>
- (spc_matches_prefix _s1922_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3380_ (_s3381_ : string) : M (option ((mword 4 * mword 4))) :=
+ let _s3382_ := _s3381_ in
+ (if string_startswith _s3382_ "fence" then
+ (match (string_drop _s3382_ (projT1 (string_length "fence"))) with
+ | _s3383_ =>
+ (spc_matches_prefix _s3383_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1923_ _)) =>
- (match (string_drop _s1922_ _s1923_) with
- | _s1924_ =>
- (fence_bits_matches_prefix _s1924_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3384_ _) =>
+ (match (string_drop _s3383_ _s3384_) with
+ | _s3385_ =>
+ (fence_bits_matches_prefix _s3385_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((pred, existT _ _s1925_ _)) =>
- (match (string_drop _s1924_ _s1925_) with
- | _s1926_ =>
- (sep_matches_prefix _s1926_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (pred, existT _ _s3386_ _) =>
+ (match (string_drop _s3385_ _s3386_) with
+ | _s3387_ =>
+ (sep_matches_prefix _s3387_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1927_ _)) =>
- (match (string_drop _s1926_ _s1927_) with
- | _s1928_ =>
- (fence_bits_matches_prefix _s1928_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3388_ _) =>
+ (match (string_drop _s3387_ _s3388_) with
+ | _s3389_ =>
+ (fence_bits_matches_prefix _s3389_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((succ, existT _ _s1929_ _)) =>
- let p0_ := string_drop _s1928_ _s1929_ in
- if ((generic_eq p0_ "")) then Some ((pred, succ))
- else None
- | _ => None
- end)
- : option ((mword 4 * mword 4)))
+ returnm (match w__3 with
+ | Some (succ, existT _ _s3390_ _) =>
+ let p0_ := string_drop _s3389_ _s3390_ in
+ if generic_eq p0_ "" then Some (pred, succ)
+ else None
+ | _ => None
+ end)
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- | _ => returnm (None : option ((mword 4 * mword 4)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4)))
end)
: M (option ((mword 4 * mword 4)))
- else returnm (None : option ((mword 4 * mword 4))))
+ else returnm None)
: M (option ((mword 4 * mword 4))).
-Definition _s1902_ (_s1903_ : string)
-: M (option ((sopw * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1903_ with
- | _s1904_ =>
- (shiftiwop_mnemonic_matches_prefix _s1904_) >>= fun w__0 : option ((sopw * {n : Z & ArithFact (n >=
+Definition _s3363_ (_s3364_ : string) : M (option ((sopw * mword 5 * mword 5 * mword 5))) :=
+ (match _s3364_ with
+ | _s3365_ =>
+ (shiftiwop_mnemonic_matches_prefix _s3365_) >>= fun w__0 : option ((sopw * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1905_ _)) =>
- (match (string_drop _s1904_ _s1905_) with
- | _s1906_ =>
- (spc_matches_prefix _s1906_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3366_ _) =>
+ (match (string_drop _s3365_ _s3366_) with
+ | _s3367_ =>
+ (spc_matches_prefix _s3367_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1907_ _)) =>
- (match (string_drop _s1906_ _s1907_) with
- | _s1908_ =>
- (reg_name_matches_prefix _s1908_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3368_ _) =>
+ (match (string_drop _s3367_ _s3368_) with
+ | _s3369_ =>
+ (reg_name_matches_prefix _s3369_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1909_ _)) =>
- (match (string_drop _s1908_ _s1909_) with
- | _s1910_ =>
- (sep_matches_prefix _s1910_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3370_ _) =>
+ (match (string_drop _s3369_ _s3370_) with
+ | _s3371_ =>
+ (sep_matches_prefix _s3371_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1911_ _)) =>
- (match (string_drop _s1910_ _s1911_) with
- | _s1912_ =>
- (reg_name_matches_prefix _s1912_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3372_ _) =>
+ (match (string_drop _s3371_ _s3372_) with
+ | _s3373_ =>
+ (reg_name_matches_prefix _s3373_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1913_ _)) =>
- (match (string_drop _s1912_ _s1913_) with
- | _s1914_ =>
- (sep_matches_prefix _s1914_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3374_ _) =>
+ (match (string_drop _s3373_ _s3374_) with
+ | _s3375_ =>
+ (sep_matches_prefix _s3375_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s1915_ _)) =>
- match (string_drop _s1914_ _s1915_) with
- | _s1916_ =>
- match (hex_bits_5_matches_prefix
- _s1916_) with
- | Some ((shamt, existT _ _s1917_ _)) =>
- let p0_ :=
- string_drop _s1916_ _s1917_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sopw * mword 5 * mword 5 * mword 5)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s3376_ _) =>
+ match (string_drop _s3375_ _s3376_) with
+ | _s3377_ =>
+ match (hex_bits_5_matches_prefix
+ _s3377_) with
+ | Some (shamt, existT _ _s3378_ _) =>
+ let p0_ :=
+ string_drop _s3377_ _s3378_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sopw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5))).
-Definition _s1885_ (_s1886_ : string)
-: M (option ((ropw * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1886_ with
- | _s1887_ =>
- (rtypew_mnemonic_matches_prefix _s1887_) >>= fun w__0 : option ((ropw * {n : Z & ArithFact (n >=
+Definition _s3346_ (_s3347_ : string) : M (option ((ropw * mword 5 * mword 5 * mword 5))) :=
+ (match _s3347_ with
+ | _s3348_ =>
+ (rtypew_mnemonic_matches_prefix _s3348_) >>= fun w__0 : option ((ropw * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1888_ _)) =>
- (match (string_drop _s1887_ _s1888_) with
- | _s1889_ =>
- (spc_matches_prefix _s1889_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3349_ _) =>
+ (match (string_drop _s3348_ _s3349_) with
+ | _s3350_ =>
+ (spc_matches_prefix _s3350_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1890_ _)) =>
- (match (string_drop _s1889_ _s1890_) with
- | _s1891_ =>
- (reg_name_matches_prefix _s1891_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3351_ _) =>
+ (match (string_drop _s3350_ _s3351_) with
+ | _s3352_ =>
+ (reg_name_matches_prefix _s3352_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1892_ _)) =>
- (match (string_drop _s1891_ _s1892_) with
- | _s1893_ =>
- (sep_matches_prefix _s1893_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3353_ _) =>
+ (match (string_drop _s3352_ _s3353_) with
+ | _s3354_ =>
+ (sep_matches_prefix _s3354_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1894_ _)) =>
- (match (string_drop _s1893_ _s1894_) with
- | _s1895_ =>
- (reg_name_matches_prefix _s1895_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3355_ _) =>
+ (match (string_drop _s3354_ _s3355_) with
+ | _s3356_ =>
+ (reg_name_matches_prefix _s3356_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1896_ _)) =>
- (match (string_drop _s1895_ _s1896_) with
- | _s1897_ =>
- (sep_matches_prefix _s1897_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3357_ _) =>
+ (match (string_drop _s3356_ _s3357_) with
+ | _s3358_ =>
+ (sep_matches_prefix _s3358_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1898_ _)) =>
- (match (string_drop _s1897_ _s1898_) with
- | _s1899_ =>
- (reg_name_matches_prefix _s1899_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3359_ _) =>
+ (match (string_drop _s3358_ _s3359_) with
+ | _s3360_ =>
+ (reg_name_matches_prefix _s3360_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s1900_ _)) =>
- let p0_ :=
- string_drop _s1899_ _s1900_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((ropw * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s3361_ _) =>
+ let p0_ :=
+ string_drop _s3360_ _s3361_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((ropw * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5))).
-Definition _s1868_ (_s1869_ : string)
-: M (option ((sop * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1869_ with
- | _s1870_ =>
- (shiftw_mnemonic_matches_prefix _s1870_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=
+Definition _s3329_ (_s3330_ : string) : M (option ((sop * mword 5 * mword 5 * mword 5))) :=
+ (match _s3330_ with
+ | _s3331_ =>
+ (shiftw_mnemonic_matches_prefix _s3331_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1871_ _)) =>
- (match (string_drop _s1870_ _s1871_) with
- | _s1872_ =>
- (spc_matches_prefix _s1872_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3332_ _) =>
+ (match (string_drop _s3331_ _s3332_) with
+ | _s3333_ =>
+ (spc_matches_prefix _s3333_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1873_ _)) =>
- (match (string_drop _s1872_ _s1873_) with
- | _s1874_ =>
- (reg_name_matches_prefix _s1874_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3334_ _) =>
+ (match (string_drop _s3333_ _s3334_) with
+ | _s3335_ =>
+ (reg_name_matches_prefix _s3335_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1875_ _)) =>
- (match (string_drop _s1874_ _s1875_) with
- | _s1876_ =>
- (sep_matches_prefix _s1876_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3336_ _) =>
+ (match (string_drop _s3335_ _s3336_) with
+ | _s3337_ =>
+ (sep_matches_prefix _s3337_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1877_ _)) =>
- (match (string_drop _s1876_ _s1877_) with
- | _s1878_ =>
- (reg_name_matches_prefix _s1878_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3338_ _) =>
+ (match (string_drop _s3337_ _s3338_) with
+ | _s3339_ =>
+ (reg_name_matches_prefix _s3339_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1879_ _)) =>
- (match (string_drop _s1878_ _s1879_) with
- | _s1880_ =>
- (sep_matches_prefix _s1880_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3340_ _) =>
+ (match (string_drop _s3339_ _s3340_) with
+ | _s3341_ =>
+ (sep_matches_prefix _s3341_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s1881_ _)) =>
- match (string_drop _s1880_ _s1881_) with
- | _s1882_ =>
- match (hex_bits_5_matches_prefix
- _s1882_) with
- | Some ((shamt, existT _ _s1883_ _)) =>
- let p0_ :=
- string_drop _s1882_ _s1883_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sop * mword 5 * mword 5 * mword 5)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s3342_ _) =>
+ match (string_drop _s3341_ _s3342_) with
+ | _s3343_ =>
+ match (hex_bits_5_matches_prefix
+ _s3343_) with
+ | Some (shamt, existT _ _s3344_ _) =>
+ let p0_ :=
+ string_drop _s3343_ _s3344_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5))).
-Definition _s1852_ (_s1853_ : string)
-: M (option ((mword 5 * mword 5 * mword 12))) :=
-
- let _s1854_ := _s1853_ in
- (if ((string_startswith _s1854_ "addiw")) then
- (match (string_drop _s1854_ (projT1 (string_length "addiw"))) with
- | _s1855_ =>
- (spc_matches_prefix _s1855_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3313_ (_s3314_ : string) : M (option ((mword 5 * mword 5 * mword 12))) :=
+ let _s3315_ := _s3314_ in
+ (if string_startswith _s3315_ "addiw" then
+ (match (string_drop _s3315_ (projT1 (string_length "addiw"))) with
+ | _s3316_ =>
+ (spc_matches_prefix _s3316_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1856_ _)) =>
- (match (string_drop _s1855_ _s1856_) with
- | _s1857_ =>
- (reg_name_matches_prefix _s1857_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3317_ _) =>
+ (match (string_drop _s3316_ _s3317_) with
+ | _s3318_ =>
+ (reg_name_matches_prefix _s3318_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1858_ _)) =>
- (match (string_drop _s1857_ _s1858_) with
- | _s1859_ =>
- (sep_matches_prefix _s1859_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3319_ _) =>
+ (match (string_drop _s3318_ _s3319_) with
+ | _s3320_ =>
+ (sep_matches_prefix _s3320_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1860_ _)) =>
- (match (string_drop _s1859_ _s1860_) with
- | _s1861_ =>
- (reg_name_matches_prefix _s1861_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3321_ _) =>
+ (match (string_drop _s3320_ _s3321_) with
+ | _s3322_ =>
+ (reg_name_matches_prefix _s3322_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s1862_ _)) =>
- (match (string_drop _s1861_ _s1862_) with
- | _s1863_ =>
- (sep_matches_prefix _s1863_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3323_ _) =>
+ (match (string_drop _s3322_ _s3323_) with
+ | _s3324_ =>
+ (sep_matches_prefix _s3324_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s1864_ _)) =>
- match (string_drop _s1863_ _s1864_) with
- | _s1865_ =>
- match (hex_bits_12_matches_prefix _s1865_) with
- | Some ((imm, existT _ _s1866_ _)) =>
- let p0_ := string_drop _s1865_ _s1866_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rd, rs1, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 12)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s3325_ _) =>
+ match (string_drop _s3324_ _s3325_) with
+ | _s3326_ =>
+ match (hex_bits_12_matches_prefix _s3326_) with
+ | Some (imm, existT _ _s3327_ _) =>
+ let p0_ := string_drop _s3326_ _s3327_ in
+ if generic_eq p0_ "" then
+ Some (rd, rs1, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- else returnm (None : option ((mword 5 * mword 5 * mword 12))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 12))).
-Definition _s1824_ (_s1825_ : string)
+Definition _s3285_ (_s3286_ : string)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5))) :=
-
- let _s1826_ := _s1825_ in
- (if ((string_startswith _s1826_ "s")) then
- (match (string_drop _s1826_ (projT1 (string_length "s"))) with
- | _s1827_ =>
- (size_mnemonic_matches_prefix _s1827_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s3287_ := _s3286_ in
+ (if string_startswith _s3287_ "s" then
+ (match (string_drop _s3287_ (projT1 (string_length "s"))) with
+ | _s3288_ =>
+ (size_mnemonic_matches_prefix _s3288_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s1828_ _)) =>
- (match (string_drop _s1827_ _s1828_) with
- | _s1829_ =>
- (maybe_aq_matches_prefix _s1829_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s3289_ _) =>
+ (match (string_drop _s3288_ _s3289_) with
+ | _s3290_ =>
+ (maybe_aq_matches_prefix _s3290_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s1830_ _)) =>
- (match (string_drop _s1829_ _s1830_) with
- | _s1831_ =>
- (maybe_rl_matches_prefix _s1831_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s3291_ _) =>
+ (match (string_drop _s3290_ _s3291_) with
+ | _s3292_ =>
+ (maybe_rl_matches_prefix _s3292_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s1832_ _)) =>
- (match (string_drop _s1831_ _s1832_) with
- | _s1833_ =>
- (spc_matches_prefix _s1833_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s3293_ _) =>
+ (match (string_drop _s3292_ _s3293_) with
+ | _s3294_ =>
+ (spc_matches_prefix _s3294_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1834_ _)) =>
- (match (string_drop _s1833_ _s1834_) with
- | _s1835_ =>
- (reg_name_matches_prefix _s1835_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3295_ _) =>
+ (match (string_drop _s3294_ _s3295_) with
+ | _s3296_ =>
+ (reg_name_matches_prefix _s3296_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs2, existT _ _s1836_ _)) =>
- (match (string_drop _s1835_ _s1836_) with
- | _s1837_ =>
- (sep_matches_prefix _s1837_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s3297_ _) =>
+ (match (string_drop _s3296_ _s3297_) with
+ | _s3298_ =>
+ (sep_matches_prefix _s3298_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1838_ _)) =>
- (match (string_drop _s1837_ _s1838_) with
- | _s1839_ =>
- (match (hex_bits_12_matches_prefix _s1839_) with
- | Some ((imm, existT _ _s1840_ _)) =>
- (match (string_drop _s1839_ _s1840_) with
- | _s1841_ =>
- (opt_spc_matches_prefix _s1841_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3299_ _) =>
+ (match (string_drop _s3298_ _s3299_) with
+ | _s3300_ =>
+ (match (hex_bits_12_matches_prefix _s3300_) with
+ | Some (imm, existT _ _s3301_ _) =>
+ (match (string_drop _s3300_ _s3301_) with
+ | _s3302_ =>
+ (opt_spc_matches_prefix _s3302_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s1842_ _)) =>
- let _s1843_ :=
- string_drop _s1841_ _s1842_ in
- (if ((string_startswith
- _s1843_ "(")) then
- (match (string_drop _s1843_
+ | Some (tt, existT _ _s3303_ _) =>
+ let _s3304_ :=
+ string_drop _s3302_ _s3303_ in
+ (if string_startswith _s3304_
+ "(" then
+ (match (string_drop _s3304_
(projT1
(string_length
"("))) with
- | _s1844_ =>
+ | _s3305_ =>
(opt_spc_matches_prefix
- _s1844_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ _s3305_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
| Some
- ((tt, existT _ _s1845_ _)) =>
+ (tt, existT _ _s3306_ _) =>
(match (string_drop
- _s1844_
- _s1845_) with
- | _s1846_ =>
+ _s3305_
+ _s3306_) with
+ | _s3307_ =>
(reg_name_matches_prefix
- _s1846_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s3307_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((rs1, existT _ _s1847_ _)) =>
+ (rs1, existT _ _s3308_ _) =>
(match (string_drop
- _s1846_
- _s1847_) with
- | _s1848_ =>
+ _s3307_
+ _s3308_) with
+ | _s3309_ =>
(opt_spc_matches_prefix
- _s1848_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=
+ _s3309_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__9 with
- | Some
- ((tt, existT _ _s1849_ _)) =>
- let _s1850_ :=
- string_drop
- _s1848_
- _s1849_ in
- if
- ((string_startswith
- _s1850_
- ")"))
- then
- let p0_ :=
- string_drop
- _s1850_
- (projT1
- (string_length
- ")")) in
- if
- ((generic_eq
- p0_
- ""))
- then
- Some
- ((size, aq, rl, rs2, imm, rs1))
- else
- None
- else
- None
- | _ =>
- None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm (match w__9 with
+ | Some
+ (tt, existT _ _s3310_ _) =>
+ let _s3311_ :=
+ string_drop
+ _s3309_
+ _s3310_ in
+ if
+ string_startswith
+ _s3311_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s3311_
+ (projT1
+ (string_length
+ ")")) in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (size, aq, rl, rs2, imm, rs1)
+ else
+ None
+ else
+ None
+ | _ =>
+ None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- else
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)))
- else returnm (None : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5))).
-Definition _s1794_ (_s1795_ : string)
+Definition _s3255_ (_s3256_ : string)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5))) :=
-
- let _s1796_ := _s1795_ in
- (if ((string_startswith _s1796_ "l")) then
- (match (string_drop _s1796_ (projT1 (string_length "l"))) with
- | _s1797_ =>
- (size_mnemonic_matches_prefix _s1797_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s3257_ := _s3256_ in
+ (if string_startswith _s3257_ "l" then
+ (match (string_drop _s3257_ (projT1 (string_length "l"))) with
+ | _s3258_ =>
+ (size_mnemonic_matches_prefix _s3258_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s1798_ _)) =>
- (match (string_drop _s1797_ _s1798_) with
- | _s1799_ =>
- (maybe_u_matches_prefix _s1799_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s3259_ _) =>
+ (match (string_drop _s3258_ _s3259_) with
+ | _s3260_ =>
+ (maybe_u_matches_prefix _s3260_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((is_unsigned, existT _ _s1800_ _)) =>
- (match (string_drop _s1799_ _s1800_) with
- | _s1801_ =>
- (maybe_aq_matches_prefix _s1801_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (is_unsigned, existT _ _s3261_ _) =>
+ (match (string_drop _s3260_ _s3261_) with
+ | _s3262_ =>
+ (maybe_aq_matches_prefix _s3262_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((aq, existT _ _s1802_ _)) =>
- (match (string_drop _s1801_ _s1802_) with
- | _s1803_ =>
- (maybe_rl_matches_prefix _s1803_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s3263_ _) =>
+ (match (string_drop _s3262_ _s3263_) with
+ | _s3264_ =>
+ (maybe_rl_matches_prefix _s3264_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rl, existT _ _s1804_ _)) =>
- (match (string_drop _s1803_ _s1804_) with
- | _s1805_ =>
- (spc_matches_prefix _s1805_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s3265_ _) =>
+ (match (string_drop _s3264_ _s3265_) with
+ | _s3266_ =>
+ (spc_matches_prefix _s3266_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s1806_ _)) =>
- (match (string_drop _s1805_ _s1806_) with
- | _s1807_ =>
- (reg_name_matches_prefix _s1807_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3267_ _) =>
+ (match (string_drop _s3266_ _s3267_) with
+ | _s3268_ =>
+ (reg_name_matches_prefix _s3268_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((rd, existT _ _s1808_ _)) =>
- (match (string_drop _s1807_ _s1808_) with
- | _s1809_ =>
- (sep_matches_prefix _s1809_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3269_ _) =>
+ (match (string_drop _s3268_ _s3269_) with
+ | _s3270_ =>
+ (sep_matches_prefix _s3270_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s1810_ _)) =>
- (match (string_drop _s1809_ _s1810_) with
- | _s1811_ =>
+ | Some (tt, existT _ _s3271_ _) =>
+ (match (string_drop _s3270_ _s3271_) with
+ | _s3272_ =>
(match (hex_bits_12_matches_prefix
- _s1811_) with
- | Some
- ((imm, existT _ _s1812_ _)) =>
- (match (string_drop _s1811_
- _s1812_) with
- | _s1813_ =>
+ _s3272_) with
+ | Some (imm, existT _ _s3273_ _) =>
+ (match (string_drop _s3272_
+ _s3273_) with
+ | _s3274_ =>
(opt_spc_matches_prefix
- _s1813_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ _s3274_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
| Some
- ((tt, existT _ _s1814_ _)) =>
- let _s1815_ :=
- string_drop _s1813_
- _s1814_ in
- (if ((string_startswith
- _s1815_ "("))
- then
+ (tt, existT _ _s3275_ _) =>
+ let _s3276_ :=
+ string_drop _s3274_
+ _s3275_ in
+ (if string_startswith
+ _s3276_ "(" then
(match (string_drop
- _s1815_
+ _s3276_
(projT1
(string_length
"("))) with
- | _s1816_ =>
+ | _s3277_ =>
(opt_spc_matches_prefix
- _s1816_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=
+ _s3277_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((tt, existT _ _s1817_ _)) =>
+ (tt, existT _ _s3278_ _) =>
(match (string_drop
- _s1816_
- _s1817_) with
- | _s1818_ =>
+ _s3277_
+ _s3278_) with
+ | _s3279_ =>
(reg_name_matches_prefix
- _s1818_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s3279_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__9 with
| Some
- ((rs1, existT _ _s1819_ _)) =>
+ (rs1, existT _ _s3280_ _) =>
(match (string_drop
- _s1818_
- _s1819_) with
- | _s1820_ =>
+ _s3279_
+ _s3280_) with
+ | _s3281_ =>
(opt_spc_matches_prefix
- _s1820_) >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=
+ _s3281_) >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__10 with
- | Some
- ((tt, existT _ _s1821_ _)) =>
- let _s1822_ :=
- string_drop
- _s1820_
- _s1821_ in
- if
- ((string_startswith
- _s1822_
- ")"))
- then
- let p0_ :=
- string_drop
- _s1822_
- (projT1
- (string_length
- ")")) in
- if
- ((generic_eq
- p0_
- ""))
- then
- Some
- ((size, is_unsigned, aq, rl, rd, imm, rs1))
- else
- None
- else
- None
- | _ =>
- None
- end)
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm (match w__10 with
+ | Some
+ (tt, existT _ _s3282_ _) =>
+ let _s3283_ :=
+ string_drop
+ _s3281_
+ _s3282_ in
+ if
+ string_startswith
+ _s3283_
+ ")"
+ then
+ let p0_ :=
+ string_drop
+ _s3283_
+ (projT1
+ (string_length
+ ")")) in
+ if
+ generic_eq
+ p0_
+ ""
+ then
+ Some
+ (size, is_unsigned, aq, rl, rd, imm, rs1)
+ else
+ None
+ else
+ None
+ | _ =>
+ None
+ end)
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- else
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)))
- else returnm (None : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5))))
+ else returnm None)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5))).
-Definition _s1777_ (_s1778_ : string)
-: M (option ((rop * mword 5 * mword 5 * mword 5))) :=
-
- (match _s1778_ with
- | _s1779_ =>
- (rtype_mnemonic_matches_prefix _s1779_) >>= fun w__0 : option ((rop * {n : Z & ArithFact (n >=
+Definition _s3238_ (_s3239_ : string) : M (option ((rop * mword 5 * mword 5 * mword 5))) :=
+ (match _s3239_ with
+ | _s3240_ =>
+ (rtype_mnemonic_matches_prefix _s3240_) >>= fun w__0 : option ((rop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1780_ _)) =>
- (match (string_drop _s1779_ _s1780_) with
- | _s1781_ =>
- (spc_matches_prefix _s1781_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3241_ _) =>
+ (match (string_drop _s3240_ _s3241_) with
+ | _s3242_ =>
+ (spc_matches_prefix _s3242_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1782_ _)) =>
- (match (string_drop _s1781_ _s1782_) with
- | _s1783_ =>
- (reg_name_matches_prefix _s1783_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3243_ _) =>
+ (match (string_drop _s3242_ _s3243_) with
+ | _s3244_ =>
+ (reg_name_matches_prefix _s3244_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1784_ _)) =>
- (match (string_drop _s1783_ _s1784_) with
- | _s1785_ =>
- (sep_matches_prefix _s1785_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3245_ _) =>
+ (match (string_drop _s3244_ _s3245_) with
+ | _s3246_ =>
+ (sep_matches_prefix _s3246_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1786_ _)) =>
- (match (string_drop _s1785_ _s1786_) with
- | _s1787_ =>
- (reg_name_matches_prefix _s1787_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3247_ _) =>
+ (match (string_drop _s3246_ _s3247_) with
+ | _s3248_ =>
+ (reg_name_matches_prefix _s3248_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1788_ _)) =>
- (match (string_drop _s1787_ _s1788_) with
- | _s1789_ =>
- (sep_matches_prefix _s1789_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3249_ _) =>
+ (match (string_drop _s3248_ _s3249_) with
+ | _s3250_ =>
+ (sep_matches_prefix _s3250_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s1790_ _)) =>
- (match (string_drop _s1789_ _s1790_) with
- | _s1791_ =>
- (reg_name_matches_prefix _s1791_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3251_ _) =>
+ (match (string_drop _s3250_ _s3251_) with
+ | _s3252_ =>
+ (reg_name_matches_prefix _s3252_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s1792_ _)) =>
- let p0_ :=
- string_drop _s1791_ _s1792_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, rs2))
- else None
- | _ => None
- end)
- : option ((rop * mword 5 * mword 5 * mword 5)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s3253_ _) =>
+ let p0_ :=
+ string_drop _s3252_ _s3253_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, rs2)
+ else None
+ | _ => None
+ end)
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None
- : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ =>
- returnm (None : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
- | _ => returnm (None : option ((rop * mword 5 * mword 5 * mword 5)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5))).
-Definition _s1760_ (_s1761_ : string)
-: M (option ((sop * mword 5 * mword 5 * mword 6))) :=
-
- (match _s1761_ with
- | _s1762_ =>
- (shiftiop_mnemonic_matches_prefix _s1762_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=
+Definition _s3221_ (_s3222_ : string) : M (option ((sop * mword 5 * mword 5 * mword 6))) :=
+ (match _s3222_ with
+ | _s3223_ =>
+ (shiftiop_mnemonic_matches_prefix _s3223_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1763_ _)) =>
- (match (string_drop _s1762_ _s1763_) with
- | _s1764_ =>
- (spc_matches_prefix _s1764_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3224_ _) =>
+ (match (string_drop _s3223_ _s3224_) with
+ | _s3225_ =>
+ (spc_matches_prefix _s3225_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1765_ _)) =>
- (match (string_drop _s1764_ _s1765_) with
- | _s1766_ =>
- (reg_name_matches_prefix _s1766_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3226_ _) =>
+ (match (string_drop _s3225_ _s3226_) with
+ | _s3227_ =>
+ (reg_name_matches_prefix _s3227_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1767_ _)) =>
- (match (string_drop _s1766_ _s1767_) with
- | _s1768_ =>
- (sep_matches_prefix _s1768_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3228_ _) =>
+ (match (string_drop _s3227_ _s3228_) with
+ | _s3229_ =>
+ (sep_matches_prefix _s3229_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1769_ _)) =>
- (match (string_drop _s1768_ _s1769_) with
- | _s1770_ =>
- (reg_name_matches_prefix _s1770_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3230_ _) =>
+ (match (string_drop _s3229_ _s3230_) with
+ | _s3231_ =>
+ (reg_name_matches_prefix _s3231_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1771_ _)) =>
- (match (string_drop _s1770_ _s1771_) with
- | _s1772_ =>
- (sep_matches_prefix _s1772_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3232_ _) =>
+ (match (string_drop _s3231_ _s3232_) with
+ | _s3233_ =>
+ (sep_matches_prefix _s3233_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s1773_ _)) =>
- match (string_drop _s1772_ _s1773_) with
- | _s1774_ =>
- match (hex_bits_6_matches_prefix
- _s1774_) with
- | Some ((shamt, existT _ _s1775_ _)) =>
- let p0_ :=
- string_drop _s1774_ _s1775_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, shamt))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sop * mword 5 * mword 5 * mword 6)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s3234_ _) =>
+ match (string_drop _s3233_ _s3234_) with
+ | _s3235_ =>
+ match (hex_bits_6_matches_prefix
+ _s3235_) with
+ | Some (shamt, existT _ _s3236_ _) =>
+ let p0_ :=
+ string_drop _s3235_ _s3236_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, shamt)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ =>
- returnm (None
- : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ =>
- returnm (None : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 6)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6))).
-Definition _s1743_ (_s1744_ : string)
-: M (option ((iop * mword 5 * mword 5 * mword 12))) :=
-
- (match _s1744_ with
- | _s1745_ =>
- (itype_mnemonic_matches_prefix _s1745_) >>= fun w__0 : option ((iop * {n : Z & ArithFact (n >=
+Definition _s3204_ (_s3205_ : string) : M (option ((iop * mword 5 * mword 5 * mword 12))) :=
+ (match _s3205_ with
+ | _s3206_ =>
+ (itype_mnemonic_matches_prefix _s3206_) >>= fun w__0 : option ((iop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1746_ _)) =>
- (match (string_drop _s1745_ _s1746_) with
- | _s1747_ =>
- (spc_matches_prefix _s1747_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3207_ _) =>
+ (match (string_drop _s3206_ _s3207_) with
+ | _s3208_ =>
+ (spc_matches_prefix _s3208_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1748_ _)) =>
- (match (string_drop _s1747_ _s1748_) with
- | _s1749_ =>
- (reg_name_matches_prefix _s1749_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3209_ _) =>
+ (match (string_drop _s3208_ _s3209_) with
+ | _s3210_ =>
+ (reg_name_matches_prefix _s3210_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1750_ _)) =>
- (match (string_drop _s1749_ _s1750_) with
- | _s1751_ =>
- (sep_matches_prefix _s1751_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3211_ _) =>
+ (match (string_drop _s3210_ _s3211_) with
+ | _s3212_ =>
+ (sep_matches_prefix _s3212_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1752_ _)) =>
- (match (string_drop _s1751_ _s1752_) with
- | _s1753_ =>
- (reg_name_matches_prefix _s1753_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3213_ _) =>
+ (match (string_drop _s3212_ _s3213_) with
+ | _s3214_ =>
+ (reg_name_matches_prefix _s3214_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s1754_ _)) =>
- (match (string_drop _s1753_ _s1754_) with
- | _s1755_ =>
- (sep_matches_prefix _s1755_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3215_ _) =>
+ (match (string_drop _s3214_ _s3215_) with
+ | _s3216_ =>
+ (sep_matches_prefix _s3216_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s1756_ _)) =>
- match (string_drop _s1755_ _s1756_) with
- | _s1757_ =>
- match (hex_bits_12_matches_prefix
- _s1757_) with
- | Some ((imm, existT _ _s1758_ _)) =>
- let p0_ :=
- string_drop _s1757_ _s1758_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rd, rs1, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((iop * mword 5 * mword 5 * mword 12)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s3217_ _) =>
+ match (string_drop _s3216_ _s3217_) with
+ | _s3218_ =>
+ match (hex_bits_12_matches_prefix
+ _s3218_) with
+ | Some (imm, existT _ _s3219_ _) =>
+ let p0_ :=
+ string_drop _s3218_ _s3219_ in
+ if generic_eq p0_ "" then
+ Some (op, rd, rs1, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None
- : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ =>
- returnm (None : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((iop * mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12))).
-Definition _s1726_ (_s1727_ : string)
-: M (option ((bop * mword 5 * mword 5 * mword 13))) :=
-
- (match _s1727_ with
- | _s1728_ =>
- (btype_mnemonic_matches_prefix _s1728_) >>= fun w__0 : option ((bop * {n : Z & ArithFact (n >=
+Definition _s3187_ (_s3188_ : string) : M (option ((bop * mword 5 * mword 5 * mword 13))) :=
+ (match _s3188_ with
+ | _s3189_ =>
+ (btype_mnemonic_matches_prefix _s3189_) >>= fun w__0 : option ((bop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1729_ _)) =>
- (match (string_drop _s1728_ _s1729_) with
- | _s1730_ =>
- (spc_matches_prefix _s1730_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3190_ _) =>
+ (match (string_drop _s3189_ _s3190_) with
+ | _s3191_ =>
+ (spc_matches_prefix _s3191_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1731_ _)) =>
- (match (string_drop _s1730_ _s1731_) with
- | _s1732_ =>
- (reg_name_matches_prefix _s1732_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3192_ _) =>
+ (match (string_drop _s3191_ _s3192_) with
+ | _s3193_ =>
+ (reg_name_matches_prefix _s3193_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rs1, existT _ _s1733_ _)) =>
- (match (string_drop _s1732_ _s1733_) with
- | _s1734_ =>
- (sep_matches_prefix _s1734_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3194_ _) =>
+ (match (string_drop _s3193_ _s3194_) with
+ | _s3195_ =>
+ (sep_matches_prefix _s3195_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s1735_ _)) =>
- (match (string_drop _s1734_ _s1735_) with
- | _s1736_ =>
- (reg_name_matches_prefix _s1736_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3196_ _) =>
+ (match (string_drop _s3195_ _s3196_) with
+ | _s3197_ =>
+ (reg_name_matches_prefix _s3197_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs2, existT _ _s1737_ _)) =>
- (match (string_drop _s1736_ _s1737_) with
- | _s1738_ =>
- (sep_matches_prefix _s1738_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s3198_ _) =>
+ (match (string_drop _s3197_ _s3198_) with
+ | _s3199_ =>
+ (sep_matches_prefix _s3199_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s1739_ _)) =>
- match (string_drop _s1738_ _s1739_) with
- | _s1740_ =>
- match (hex_bits_13_matches_prefix
- _s1740_) with
- | Some ((imm, existT _ _s1741_ _)) =>
- let p0_ :=
- string_drop _s1740_ _s1741_ in
- if ((generic_eq p0_ "")) then
- Some
- ((op, rs1, rs2, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((bop * mword 5 * mword 5 * mword 13)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s3200_ _) =>
+ match (string_drop _s3199_ _s3200_) with
+ | _s3201_ =>
+ match (hex_bits_13_matches_prefix
+ _s3201_) with
+ | Some (imm, existT _ _s3202_ _) =>
+ let p0_ :=
+ string_drop _s3201_ _s3202_ in
+ if generic_eq p0_ "" then
+ Some (op, rs1, rs2, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ =>
- returnm (None
- : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ =>
- returnm (None : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ => returnm (None : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ => returnm (None : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
- | _ => returnm (None : option ((bop * mword 5 * mword 5 * mword 13)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13))).
-Definition _s1710_ (_s1711_ : string)
-: M (option ((mword 5 * mword 5 * mword 12))) :=
-
- let _s1712_ := _s1711_ in
- (if ((string_startswith _s1712_ "jalr")) then
- (match (string_drop _s1712_ (projT1 (string_length "jalr"))) with
- | _s1713_ =>
- (spc_matches_prefix _s1713_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3171_ (_s3172_ : string) : M (option ((mword 5 * mword 5 * mword 12))) :=
+ let _s3173_ := _s3172_ in
+ (if string_startswith _s3173_ "jalr" then
+ (match (string_drop _s3173_ (projT1 (string_length "jalr"))) with
+ | _s3174_ =>
+ (spc_matches_prefix _s3174_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1714_ _)) =>
- (match (string_drop _s1713_ _s1714_) with
- | _s1715_ =>
- (reg_name_matches_prefix _s1715_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3175_ _) =>
+ (match (string_drop _s3174_ _s3175_) with
+ | _s3176_ =>
+ (reg_name_matches_prefix _s3176_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1716_ _)) =>
- (match (string_drop _s1715_ _s1716_) with
- | _s1717_ =>
- (sep_matches_prefix _s1717_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3177_ _) =>
+ (match (string_drop _s3176_ _s3177_) with
+ | _s3178_ =>
+ (sep_matches_prefix _s3178_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s1718_ _)) =>
- (match (string_drop _s1717_ _s1718_) with
- | _s1719_ =>
- (reg_name_matches_prefix _s1719_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3179_ _) =>
+ (match (string_drop _s3178_ _s3179_) with
+ | _s3180_ =>
+ (reg_name_matches_prefix _s3180_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s1720_ _)) =>
- (match (string_drop _s1719_ _s1720_) with
- | _s1721_ =>
- (sep_matches_prefix _s1721_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s3181_ _) =>
+ (match (string_drop _s3180_ _s3181_) with
+ | _s3182_ =>
+ (sep_matches_prefix _s3182_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s1722_ _)) =>
- match (string_drop _s1721_ _s1722_) with
- | _s1723_ =>
- match (hex_bits_12_matches_prefix _s1723_) with
- | Some ((imm, existT _ _s1724_ _)) =>
- let p0_ := string_drop _s1723_ _s1724_ in
- if ((generic_eq p0_ "")) then
- Some
- ((rd, rs1, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 12)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s3183_ _) =>
+ match (string_drop _s3182_ _s3183_) with
+ | _s3184_ =>
+ match (hex_bits_12_matches_prefix _s3184_) with
+ | Some (imm, existT _ _s3185_ _) =>
+ let p0_ := string_drop _s3184_ _s3185_ in
+ if generic_eq p0_ "" then
+ Some (rd, rs1, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12)))
end)
: M (option ((mword 5 * mword 5 * mword 12)))
- else returnm (None : option ((mword 5 * mword 5 * mword 12))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 12))).
-Definition _s1698_ (_s1699_ : string)
-: M (option ((mword 5 * mword 21))) :=
-
- let _s1700_ := _s1699_ in
- (if ((string_startswith _s1700_ "jal")) then
- (match (string_drop _s1700_ (projT1 (string_length "jal"))) with
- | _s1701_ =>
- (spc_matches_prefix _s1701_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s3159_ (_s3160_ : string) : M (option ((mword 5 * mword 21))) :=
+ let _s3161_ := _s3160_ in
+ (if string_startswith _s3161_ "jal" then
+ (match (string_drop _s3161_ (projT1 (string_length "jal"))) with
+ | _s3162_ =>
+ (spc_matches_prefix _s3162_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s1702_ _)) =>
- (match (string_drop _s1701_ _s1702_) with
- | _s1703_ =>
- (reg_name_matches_prefix _s1703_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3163_ _) =>
+ (match (string_drop _s3162_ _s3163_) with
+ | _s3164_ =>
+ (reg_name_matches_prefix _s3164_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s1704_ _)) =>
- (match (string_drop _s1703_ _s1704_) with
- | _s1705_ =>
- (sep_matches_prefix _s1705_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3165_ _) =>
+ (match (string_drop _s3164_ _s3165_) with
+ | _s3166_ =>
+ (sep_matches_prefix _s3166_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s1706_ _)) =>
- match (string_drop _s1705_ _s1706_) with
- | _s1707_ =>
- match (hex_bits_21_matches_prefix _s1707_) with
- | Some ((imm, existT _ _s1708_ _)) =>
- let p0_ := string_drop _s1707_ _s1708_ in
- if ((generic_eq p0_ "")) then Some ((rd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 21)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s3167_ _) =>
+ match (string_drop _s3166_ _s3167_) with
+ | _s3168_ =>
+ match (hex_bits_21_matches_prefix _s3168_) with
+ | Some (imm, existT _ _s3169_ _) =>
+ let p0_ := string_drop _s3168_ _s3169_ in
+ if generic_eq p0_ "" then Some (rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 21)))
- | _ => returnm (None : option ((mword 5 * mword 21)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 21)))
end)
: M (option ((mword 5 * mword 21)))
- | _ => returnm (None : option ((mword 5 * mword 21)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 21)))
end)
: M (option ((mword 5 * mword 21)))
- else returnm (None : option ((mword 5 * mword 21))))
+ else returnm None)
: M (option ((mword 5 * mword 21))).
-Definition _s1685_ (_s1686_ : string)
-: M (option ((uop * mword 5 * mword 20))) :=
-
- (match _s1686_ with
- | _s1687_ =>
- (utype_mnemonic_matches_prefix _s1687_) >>= fun w__0 : option ((uop * {n : Z & ArithFact (n >=
+Definition _s3146_ (_s3147_ : string) : M (option ((uop * mword 5 * mword 20))) :=
+ (match _s3147_ with
+ | _s3148_ =>
+ (utype_mnemonic_matches_prefix _s3148_) >>= fun w__0 : option ((uop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s1688_ _)) =>
- (match (string_drop _s1687_ _s1688_) with
- | _s1689_ =>
- (spc_matches_prefix _s1689_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s3149_ _) =>
+ (match (string_drop _s3148_ _s3149_) with
+ | _s3150_ =>
+ (spc_matches_prefix _s3150_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s1690_ _)) =>
- (match (string_drop _s1689_ _s1690_) with
- | _s1691_ =>
- (reg_name_matches_prefix _s1691_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s3151_ _) =>
+ (match (string_drop _s3150_ _s3151_) with
+ | _s3152_ =>
+ (reg_name_matches_prefix _s3152_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s1692_ _)) =>
- (match (string_drop _s1691_ _s1692_) with
- | _s1693_ =>
- (sep_matches_prefix _s1693_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s3153_ _) =>
+ (match (string_drop _s3152_ _s3153_) with
+ | _s3154_ =>
+ (sep_matches_prefix _s3154_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((tt, existT _ _s1694_ _)) =>
- match (string_drop _s1693_ _s1694_) with
- | _s1695_ =>
- match (hex_bits_20_matches_prefix _s1695_) with
- | Some ((imm, existT _ _s1696_ _)) =>
- let p0_ := string_drop _s1695_ _s1696_ in
- if ((generic_eq p0_ "")) then Some ((op, rd, imm))
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((uop * mword 5 * mword 20)))
+ returnm (match w__3 with
+ | Some (tt, existT _ _s3155_ _) =>
+ match (string_drop _s3154_ _s3155_) with
+ | _s3156_ =>
+ match (hex_bits_20_matches_prefix _s3156_) with
+ | Some (imm, existT _ _s3157_ _) =>
+ let p0_ := string_drop _s3156_ _s3157_ in
+ if generic_eq p0_ "" then Some (op, rd, imm)
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((uop * mword 5 * mword 20)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20)))
end)
: M (option ((uop * mword 5 * mword 20)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20)))
end)
: M (option ((uop * mword 5 * mword 20)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20)))
end)
: M (option ((uop * mword 5 * mword 20))).
-Definition assembly_backwards_matches (arg_ : string)
-: M (bool) :=
-
- let _s1697_ := arg_ in
- (_s1685_ _s1697_) >>= fun w__0 : option ((uop * mword 5 * mword 20)) =>
- (if ((match w__0 with | Some ((op, rd, imm)) => true | _ => false end)) then
- (_s1685_ _s1697_) >>= fun w__1 : option ((uop * mword 5 * mword 20)) =>
- (match w__1 with
- | Some ((op, rd, imm)) =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | _ => exit tt : M (bool)
- end)
+Definition assembly_backwards_matches (arg_ : string) : M (bool) :=
+ let _s3158_ := arg_ in
+ (_s3146_ _s3158_) >>= fun w__0 : option ((uop * mword 5 * mword 20)) =>
+ (if match w__0 with | Some (op, rd, imm) => true | _ => false end then
+ (_s3146_ _s3158_) >>= fun w__1 : option ((uop * mword 5 * mword 20)) =>
+ (match w__1 with | Some (op, rd, imm) => returnm true | _ => exit tt : M (bool) end)
: M (bool)
else
- (_s1698_ _s1697_) >>= fun w__4 : option ((mword 5 * mword 21)) =>
- (if ((match w__4 with | Some ((rd, imm)) => true | _ => false end)) then
- (_s1698_ _s1697_) >>= fun w__5 : option ((mword 5 * mword 21)) =>
- (match w__5 with
- | Some ((rd, imm)) =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | _ => exit tt : M (bool)
- end)
+ (_s3159_ _s3158_) >>= fun w__4 : option ((mword 5 * mword 21)) =>
+ (if match w__4 with | Some (rd, imm) => true | _ => false end then
+ (_s3159_ _s3158_) >>= fun w__5 : option ((mword 5 * mword 21)) =>
+ (match w__5 with | Some (rd, imm) => returnm true | _ => exit tt : M (bool) end)
: M (bool)
else
- (_s1710_ _s1697_) >>= fun w__8 : option ((mword 5 * mword 5 * mword 12)) =>
- (if ((match w__8 with | Some ((rd, rs1, imm)) => true | _ => false end)) then
- (_s1710_ _s1697_) >>= fun w__9 : option ((mword 5 * mword 5 * mword 12)) =>
- (match w__9 with
- | Some ((rd, rs1, imm)) =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
- | _ => exit tt : M (bool)
- end)
+ (_s3171_ _s3158_) >>= fun w__8 : option ((mword 5 * mword 5 * mword 12)) =>
+ (if match w__8 with | Some (rd, rs1, imm) => true | _ => false end then
+ (_s3171_ _s3158_) >>= fun w__9 : option ((mword 5 * mword 5 * mword 12)) =>
+ (match w__9 with | Some (rd, rs1, imm) => returnm true | _ => exit tt : M (bool) end)
: M (bool)
else
- (_s1726_ _s1697_) >>= fun w__12 : option ((bop * mword 5 * mword 5 * mword 13)) =>
- (if ((match w__12 with | Some ((op, rs1, rs2, imm)) => true | _ => false end)) then
- (_s1726_ _s1697_) >>= fun w__13 : option ((bop * mword 5 * mword 5 * mword 13)) =>
+ (_s3187_ _s3158_) >>= fun w__12 : option ((bop * mword 5 * mword 5 * mword 13)) =>
+ (if match w__12 with | Some (op, rs1, rs2, imm) => true | _ => false end then
+ (_s3187_ _s3158_) >>= fun w__13 : option ((bop * mword 5 * mword 5 * mword 13)) =>
(match w__13 with
- | Some ((op, rs1, rs2, imm)) =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (op, rs1, rs2, imm) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1743_ _s1697_) >>= fun w__16 : option ((iop * mword 5 * mword 5 * mword 12)) =>
- (if ((match w__16 with | Some ((op, rd, rs1, imm)) => true | _ => false end)) then
- (_s1743_ _s1697_) >>= fun w__17 : option ((iop * mword 5 * mword 5 * mword 12)) =>
+ (_s3204_ _s3158_) >>= fun w__16 : option ((iop * mword 5 * mword 5 * mword 12)) =>
+ (if match w__16 with | Some (op, rd, rs1, imm) => true | _ => false end then
+ (_s3204_ _s3158_) >>= fun w__17 : option ((iop * mword 5 * mword 5 * mword 12)) =>
(match w__17 with
- | Some ((op, rd, rs1, imm)) =>
- returnm (projT1 (build_ex true : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (op, rd, rs1, imm) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1760_ _s1697_) >>= fun w__20 : option ((sop * mword 5 * mword 5 * mword 6)) =>
- (if ((match w__20 with | Some ((op, rd, rs1, shamt)) => true | _ => false end))
- then
- (_s1760_ _s1697_) >>= fun w__21 : option ((sop * mword 5 * mword 5 * mword 6)) =>
+ (_s3221_ _s3158_) >>= fun w__20 : option ((sop * mword 5 * mword 5 * mword 6)) =>
+ (if match w__20 with | Some (op, rd, rs1, shamt) => true | _ => false end then
+ (_s3221_ _s3158_) >>= fun w__21 : option ((sop * mword 5 * mword 5 * mword 6)) =>
(match w__21 with
- | Some ((op, rd, rs1, shamt)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (op, rd, rs1, shamt) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1777_ _s1697_) >>= fun w__24 : option ((rop * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__24 with | Some ((op, rd, rs1, rs2)) => true | _ => false end))
- then
- (_s1777_ _s1697_) >>= fun w__25 : option ((rop * mword 5 * mword 5 * mword 5)) =>
+ (_s3238_ _s3158_) >>= fun w__24 : option ((rop * mword 5 * mword 5 * mword 5)) =>
+ (if match w__24 with | Some (op, rd, rs1, rs2) => true | _ => false end then
+ (_s3238_ _s3158_) >>= fun w__25 : option ((rop * mword 5 * mword 5 * mword 5)) =>
(match w__25 with
- | Some ((op, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (op, rd, rs1, rs2) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1794_ _s1697_) >>= fun w__28 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)) =>
- (if ((match w__28 with
- | Some ((size, is_unsigned, aq, rl, rd, imm, rs1)) => true
- | _ => false
- end)) then
- (_s1794_ _s1697_) >>= fun w__29 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)) =>
+ (_s3255_ _s3158_) >>= fun w__28 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)) =>
+ (if match w__28 with
+ | Some (size, is_unsigned, aq, rl, rd, imm, rs1) => true
+ | _ => false
+ end then
+ (_s3255_ _s3158_) >>= fun w__29 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5)) =>
(match w__29 with
- | Some ((size, is_unsigned, aq, rl, rd, imm, rs1)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (size, is_unsigned, aq, rl, rd, imm, rs1) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1824_ _s1697_) >>= fun w__32 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)) =>
- (if ((match w__32 with
- | Some ((size, aq, rl, rs2, imm, rs1)) => true
- | _ => false
- end)) then
- (_s1824_ _s1697_) >>= fun w__33 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)) =>
+ (_s3285_ _s3158_) >>= fun w__32 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)) =>
+ (if match w__32 with
+ | Some (size, aq, rl, rs2, imm, rs1) => true
+ | _ => false
+ end then
+ (_s3285_ _s3158_) >>= fun w__33 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5)) =>
(match w__33 with
- | Some ((size, aq, rl, rs2, imm, rs1)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (size, aq, rl, rs2, imm, rs1) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1852_ _s1697_) >>= fun w__36 : option ((mword 5 * mword 5 * mword 12)) =>
- (if ((match w__36 with
- | Some ((rd, rs1, imm)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s1852_ _s1697_) >>= fun w__37 : option ((mword 5 * mword 5 * mword 12)) =>
+ (_s3313_ _s3158_) >>= fun w__36 : option ((mword 5 * mword 5 * mword 12)) =>
+ (if match w__36 with
+ | Some (rd, rs1, imm) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s3313_ _s3158_) >>= fun w__37 : option ((mword 5 * mword 5 * mword 12)) =>
(match w__37 with
- | Some ((rd, rs1, imm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (rd, rs1, imm) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1868_ _s1697_) >>= fun w__40 : option ((sop * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__40 with
- | Some ((op, rd, rs1, shamt)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s1868_ _s1697_) >>= fun w__41 : option ((sop * mword 5 * mword 5 * mword 5)) =>
+ (_s3329_ _s3158_) >>= fun w__40 : option ((sop * mword 5 * mword 5 * mword 5)) =>
+ (if match w__40 with
+ | Some (op, rd, rs1, shamt) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s3329_ _s3158_) >>= fun w__41 : option ((sop * mword 5 * mword 5 * mword 5)) =>
(match w__41 with
- | Some ((op, rd, rs1, shamt)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (op, rd, rs1, shamt) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1885_ _s1697_) >>= fun w__44 : option ((ropw * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__44 with
- | Some ((op, rd, rs1, rs2)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s1885_ _s1697_) >>= fun w__45 : option ((ropw * mword 5 * mword 5 * mword 5)) =>
+ (_s3346_ _s3158_) >>= fun w__44 : option ((ropw * mword 5 * mword 5 * mword 5)) =>
+ (if match w__44 with
+ | Some (op, rd, rs1, rs2) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s3346_ _s3158_) >>= fun w__45 : option ((ropw * mword 5 * mword 5 * mword 5)) =>
(match w__45 with
- | Some ((op, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (op, rd, rs1, rs2) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1902_ _s1697_) >>= fun w__48 : option ((sopw * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__48 with
- | Some ((op, rd, rs1, shamt)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s1902_ _s1697_) >>= fun w__49 : option ((sopw * mword 5 * mword 5 * mword 5)) =>
+ (_s3363_ _s3158_) >>= fun w__48 : option ((sopw * mword 5 * mword 5 * mword 5)) =>
+ (if match w__48 with
+ | Some (op, rd, rs1, shamt) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s3363_ _s3158_) >>= fun w__49 : option ((sopw * mword 5 * mword 5 * mword 5)) =>
(match w__49 with
- | Some ((op, rd, rs1, shamt)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (op, rd, rs1, shamt) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1919_ _s1697_) >>= fun w__52 : option ((mword 4 * mword 4)) =>
- (if ((match w__52 with
- | Some ((pred, succ)) => true
- | _ => false
- end)) then
- (_s1919_ _s1697_) >>= fun w__53 : option ((mword 4 * mword 4)) =>
+ (_s3380_ _s3158_) >>= fun w__52 : option ((mword 4 * mword 4)) =>
+ (if match w__52 with
+ | Some (pred, succ) => true
+ | _ => false
+ end then
+ (_s3380_ _s3158_) >>= fun w__53 : option ((mword 4 * mword 4)) =>
(match w__53 with
- | Some ((pred, succ)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (pred, succ) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1931_ _s1697_) >>= fun w__56 : option ((mword 4 * mword 4)) =>
- (if ((match w__56 with
- | Some ((pred, succ)) => true
- | _ => false
- end)) then
- (_s1931_ _s1697_) >>= fun w__57 : option ((mword 4 * mword 4)) =>
+ (_s3392_ _s3158_) >>= fun w__56 : option ((mword 4 * mword 4)) =>
+ (if match w__56 with
+ | Some (pred, succ) => true
+ | _ => false
+ end then
+ (_s3392_ _s3158_) >>= fun w__57 : option ((mword 4 * mword 4)) =>
(match w__57 with
- | Some ((pred, succ)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (pred, succ) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
- else if ((generic_eq _s1697_ "fence.i")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq _s1697_ "ecall")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq _s1697_ "mret")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq _s1697_ "sret")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq _s1697_ "ebreak")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
- else if ((generic_eq _s1697_ "wfi")) then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ else if generic_eq _s3158_ "fence.i" then returnm true
+ else if generic_eq _s3158_ "ecall" then returnm true
+ else if generic_eq _s3158_ "mret" then returnm true
+ else if generic_eq _s3158_ "sret" then returnm true
+ else if generic_eq _s3158_ "ebreak" then returnm true
+ else if generic_eq _s3158_ "wfi" then returnm true
else
- (_s1943_ _s1697_) >>= fun w__60 : option ((mword 5 * mword 5)) =>
- (if ((match w__60 with
- | Some ((rs1, rs2)) => true
- | _ => false
- end)) then
- (_s1943_ _s1697_) >>= fun w__61 : option ((mword 5 * mword 5)) =>
+ (_s3404_ _s3158_) >>= fun w__60 : option ((mword 5 * mword 5)) =>
+ (if match w__60 with
+ | Some (rs1, rs2) => true
+ | _ => false
+ end then
+ (_s3404_ _s3158_) >>= fun w__61 : option ((mword 5 * mword 5)) =>
(match w__61 with
- | Some ((rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (rs1, rs2) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1955_ _s1697_) >>= fun w__64 : option ((word_width * bool * bool * mword 5 * mword 5)) =>
- (if ((match w__64 with
- | Some ((size, aq, rl, rd, rs1)) => true
- | _ => false
- end)) then
- (_s1955_ _s1697_) >>= fun w__65 : option ((word_width * bool * bool * mword 5 * mword 5)) =>
+ (_s3416_ _s3158_) >>= fun w__64 : option ((word_width * bool * bool * mword 5 * mword 5)) =>
+ (if match w__64 with
+ | Some (size, aq, rl, rd, rs1) => true
+ | _ => false
+ end then
+ (_s3416_ _s3158_) >>= fun w__65 : option ((word_width * bool * bool * mword 5 * mword 5)) =>
(match w__65 with
- | Some ((size, aq, rl, rd, rs1)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool = true)}))
+ | Some (size, aq, rl, rd, rs1) =>
+ returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1973_ _s1697_) >>= fun w__68 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__68 with
- | Some ((size, aq, rl, rd, rs1, rs2)) =>
- true
- | _ => false
- end)) then
- (_s1973_ _s1697_) >>= fun w__69 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
- (match w__69 with
- | Some ((size, aq, rl, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
+ (_s3434_ _s3158_) >>= fun w__68 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (if match w__68 with
+ | Some (size, aq, rl, rd, rs1, rs2) =>
true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ | _ => false
+ end then
+ (_s3434_ _s3158_) >>= fun w__69 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (match w__69 with
+ | Some (size, aq, rl, rd, rs1, rs2) =>
+ returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s1995_ _s1697_) >>= fun w__72 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
- (if ((match w__72 with
- | Some
- ((op, width, aq, rl, rd, rs1, rs2)) =>
- true
- | _ => false
- end)) then
- (_s1995_ _s1697_) >>= fun w__73 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
- (match w__73 with
+ (_s3456_ _s3158_) >>= fun w__72 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (if match w__72 with
| Some
- ((op, width, aq, rl, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
+ (op, width, aq, rl, rd, rs2, rs1) =>
true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ | _ => false
+ end then
+ (_s3456_ _s3158_) >>= fun w__73 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (match w__73 with
+ | Some
+ (op, width, aq, rl, rd, rs2, rs1) =>
+ returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
- else if ((generic_eq _s1697_ "c.nop"))
- then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ else if generic_eq _s3158_ "c.nop" then
+ returnm true
else
- (_s2019_ _s1697_) >>= fun w__76 : option ((mword 3 * mword 8)) =>
- (if ((match w__76 with
- | Some ((rdc, nzimm)) =>
- neq_vec nzimm
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0]
- : mword 8)
- | _ => false
- end)) then
- (_s2019_ _s1697_) >>= fun w__77 : option ((mword 3 * mword 8)) =>
+ (_s3482_ _s3158_) >>= fun w__76 : option ((mword 3 * mword 8)) =>
+ (if match w__76 with
+ | Some (rdc, nzimm) =>
+ neq_vec nzimm (Ox"00" : mword 8)
+ | _ => false
+ end then
+ (_s3482_ _s3158_) >>= fun w__77 : option ((mword 3 * mword 8)) =>
(match w__77 with
- | Some ((rdc, nzimm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ | Some (rdc, nzimm) => returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s2031_ _s1697_) >>= fun w__80 : option ((mword 3 * mword 3 * mword 5)) =>
- (if ((match w__80 with
- | Some ((rdc, rsc, uimm)) =>
- true
- | _ => false
- end)) then
- (_s2031_ _s1697_) >>= fun w__81 : option ((mword 3 * mword 3 * mword 5)) =>
+ (_s3494_ _s3158_) >>= fun w__80 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if match w__80 with
+ | Some (rdc, rsc, uimm) => true
+ | _ => false
+ end then
+ (_s3494_ _s3158_) >>= fun w__81 : option ((mword 3 * mword 3 * mword 5)) =>
(match w__81 with
- | Some ((rdc, rsc, uimm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ | Some (rdc, rsc, uimm) =>
+ returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s2047_ _s1697_) >>= fun w__84 : option ((mword 3 * mword 3 * mword 5)) =>
- (if ((match w__84 with
- | Some ((rdc, rsc, uimm)) =>
- Z.eqb 64 64
- | _ => false
- end)) then
- (_s2047_ _s1697_) >>= fun w__85 : option ((mword 3 * mword 3 * mword 5)) =>
+ (_s3510_ _s3158_) >>= fun w__84 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if match w__84 with
+ | Some (rdc, rsc, uimm) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s3510_ _s3158_) >>= fun w__85 : option ((mword 3 * mword 3 * mword 5)) =>
(match w__85 with
- | Some ((rdc, rsc, uimm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ | Some (rdc, rsc, uimm) =>
+ returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s2063_ _s1697_) >>= fun w__88 : option ((mword 3 * mword 3 * mword 5)) =>
- (if ((match w__88 with
- | Some
- ((rsc1, rsc2, uimm)) =>
- true
- | _ => false
- end)) then
- (_s2063_ _s1697_) >>= fun w__89 : option ((mword 3 * mword 3 * mword 5)) =>
- (match w__89 with
- | Some ((rsc1, rsc2, uimm)) =>
- returnm (projT1
- (build_ex
+ (_s3526_ _s3158_) >>= fun w__88 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if match w__88 with
+ | Some (rsc1, rsc2, uimm) =>
true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ | _ => false
+ end then
+ (_s3526_ _s3158_) >>= fun w__89 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__89 with
+ | Some (rsc1, rsc2, uimm) =>
+ returnm true
| _ => exit tt : M (bool)
end)
: M (bool)
else
- (_s2079_ _s1697_) >>= fun w__92 : option ((mword 3 * mword 3 * mword 5)) =>
- (if ((match w__92 with
- | Some
- ((rsc1, rsc2, uimm)) =>
- Z.eqb 64 64
- | _ => false
- end)) then
- (_s2079_ _s1697_) >>= fun w__93 : option ((mword 3 * mword 3 * mword 5)) =>
+ (_s3542_ _s3158_) >>= fun w__92 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if match w__92 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s3542_ _s3158_) >>= fun w__93 : option ((mword 3 * mword 3 * mword 5)) =>
(match w__93 with
| Some
- ((rsc1, rsc2, uimm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsc1, rsc2, uimm) =>
+ returnm true
| _ =>
exit tt : M (bool)
end)
: M (bool)
else
- (_s2095_ _s1697_) >>= fun w__96 : option ((mword 5 * mword 6)) =>
- (if ((match w__96 with
- | Some ((rsd, nzi)) =>
- andb
- (neq_vec nzi
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ => false
- end)) then
- (_s2095_ _s1697_) >>= fun w__97 : option ((mword 5 * mword 6)) =>
+ (_s3558_ _s3158_) >>= fun w__96 : option ((mword 5 * mword 6)) =>
+ (if match w__96 with
+ | Some (rsd, nzi) =>
+ andb
+ (neq_vec nzi
+ ('b"000000"
+ : mword 6))
+ (neq_vec rsd
+ zreg)
+ | _ => false
+ end then
+ (_s3558_ _s3158_) >>= fun w__97 : option ((mword 5 * mword 6)) =>
(match w__97 with
- | Some ((rsd, nzi)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ | Some (rsd, nzi) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2107_ _s1697_) >>= fun w__100 : option (mword 11) =>
- (if ((match w__100 with
- | Some (imm) =>
- Z.eqb 64 32
- | _ => false
- end)) then
- (_s2107_ _s1697_) >>= fun w__101 : option (mword 11) =>
+ (_s3570_ _s3158_) >>= fun w__100 : option (mword 11) =>
+ (if match w__100 with
+ | Some imm =>
+ Z.eqb 64 32
+ | _ => false
+ end then
+ (_s3570_ _s3158_) >>= fun w__101 : option (mword 11) =>
(match w__101 with
- | Some (imm) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ | Some imm =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2115_ _s1697_) >>= fun w__104 : option ((mword 5 * mword 6)) =>
- (if ((match w__104 with
- | Some
- ((rsd, imm)) =>
- Z.eqb 64
- 64
- | _ => false
- end)) then
- (_s2115_ _s1697_) >>= fun w__105 : option ((mword 5 * mword 6)) =>
+ (_s3578_ _s3158_) >>= fun w__104 : option ((mword 5 * mword 6)) =>
+ (if match w__104 with
+ | Some
+ (rsd, imm) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s3578_ _s3158_) >>= fun w__105 : option ((mword 5 * mword 6)) =>
(match w__105 with
| Some
- ((rsd, imm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, imm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2127_ _s1697_) >>= fun w__108 : option ((mword 5 * mword 6)) =>
- (if ((match w__108 with
- | Some
- ((rd, imm)) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end)) then
- (_s2127_
- _s1697_) >>= fun w__109 : option ((mword 5 * mword 6)) =>
+ (_s3590_ _s3158_) >>= fun w__108 : option ((mword 5 * mword 6)) =>
+ (if match w__108 with
+ | Some
+ (rd, imm) =>
+ neq_vec
+ rd zreg
+ | _ => false
+ end then
+ (_s3590_
+ _s3158_) >>= fun w__109 : option ((mword 5 * mword 6)) =>
(match w__109 with
| Some
- ((rd, imm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rd, imm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2139_
- _s1697_) >>= fun w__112 : option (mword 6) =>
- (if ((match w__112 with
- | Some
- (imm) =>
- neq_vec
- imm
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
- then
- (_s2139_
- _s1697_) >>= fun w__113 : option (mword 6) =>
+ (_s3602_
+ _s3158_) >>= fun w__112 : option (mword 6) =>
+ (if match w__112 with
+ | Some
+ imm =>
+ neq_vec
+ imm
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end then
+ (_s3602_
+ _s3158_) >>= fun w__113 : option (mword 6) =>
(match w__113 with
| Some
- (imm) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ imm =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2147_
- _s1697_) >>= fun w__116 : option ((mword 5 * mword 6)) =>
- (if ((match w__116 with
- | Some
- ((rd, imm)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- sp))))
- (neq_vec
- imm
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)))
- : bool)
- | _ =>
- false
- end))
+ (_s3610_
+ _s3158_) >>= fun w__116 : option ((mword 5 * mword 6)) =>
+ (if match w__116 with
+ | Some
+ (rd, imm) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (andb
+ (neq_vec
+ rd
+ sp)
+ (neq_vec
+ imm
+ ('b"000000"
+ : mword 6)))
+ | _ =>
+ false
+ end
then
- (_s2147_
- _s1697_) >>= fun w__117 : option ((mword 5 * mword 6)) =>
+ (_s3610_
+ _s3158_) >>= fun w__117 : option ((mword 5 * mword 6)) =>
(match w__117 with
| Some
- ((rd, imm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rd, imm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2159_
- _s1697_) >>= fun w__120 : option ((mword 3 * mword 6)) =>
- (if ((match w__120 with
- | Some
- ((rsd, shamt)) =>
- neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
+ (_s3622_
+ _s3158_) >>= fun w__120 : option ((mword 3 * mword 6)) =>
+ (if match w__120 with
+ | Some
+ (rsd, shamt) =>
+ neq_vec
+ shamt
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end
then
- (_s2159_
- _s1697_) >>= fun w__121 : option ((mword 3 * mword 6)) =>
+ (_s3622_
+ _s3158_) >>= fun w__121 : option ((mword 3 * mword 6)) =>
(match w__121 with
| Some
- ((rsd, shamt)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, shamt) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2171_
- _s1697_) >>= fun w__124 : option ((mword 3 * mword 6)) =>
- (if ((match w__124 with
- | Some
- ((rsd, shamt)) =>
- neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
+ (_s3634_
+ _s3158_) >>= fun w__124 : option ((mword 3 * mword 6)) =>
+ (if match w__124 with
+ | Some
+ (rsd, shamt) =>
+ neq_vec
+ shamt
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end
then
- (_s2171_
- _s1697_) >>= fun w__125 : option ((mword 3 * mword 6)) =>
+ (_s3634_
+ _s3158_) >>= fun w__125 : option ((mword 3 * mword 6)) =>
(match w__125 with
| Some
- ((rsd, shamt)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, shamt) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2183_
- _s1697_) >>= fun w__128 : option ((mword 3 * mword 6)) =>
+ (_s3646_
+ _s3158_) >>= fun w__128 : option ((mword 3 * mword 6)) =>
(if
- ((match w__128 with
- | Some
- ((rsd, imm)) =>
- true
- | _ =>
- false
- end))
+ match w__128 with
+ | Some
+ (rsd, imm) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2183_
- _s1697_) >>= fun w__129 : option ((mword 3 * mword 6)) =>
+ (_s3646_
+ _s3158_) >>= fun w__129 : option ((mword 3 * mword 6)) =>
(match w__129 with
| Some
- ((rsd, imm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, imm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2195_
- _s1697_) >>= fun w__132 : option ((mword 3 * mword 3)) =>
+ (_s3658_
+ _s3158_) >>= fun w__132 : option ((mword 3 * mword 3)) =>
(if
- ((match w__132 with
- | Some
- ((rsd, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__132 with
+ | Some
+ (rsd, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2195_
- _s1697_) >>= fun w__133 : option ((mword 3 * mword 3)) =>
+ (_s3658_
+ _s3158_) >>= fun w__133 : option ((mword 3 * mword 3)) =>
(match w__133 with
| Some
- ((rsd, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2207_
- _s1697_) >>= fun w__136 : option ((mword 3 * mword 3)) =>
+ (_s3670_
+ _s3158_) >>= fun w__136 : option ((mword 3 * mword 3)) =>
(if
- ((match w__136 with
- | Some
- ((rsd, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__136 with
+ | Some
+ (rsd, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2207_
- _s1697_) >>= fun w__137 : option ((mword 3 * mword 3)) =>
+ (_s3670_
+ _s3158_) >>= fun w__137 : option ((mword 3 * mword 3)) =>
(match w__137 with
| Some
- ((rsd, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2219_
- _s1697_) >>= fun w__140 : option ((mword 3 * mword 3)) =>
+ (_s3682_
+ _s3158_) >>= fun w__140 : option ((mword 3 * mword 3)) =>
(if
- ((match w__140 with
- | Some
- ((rsd, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__140 with
+ | Some
+ (rsd, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2219_
- _s1697_) >>= fun w__141 : option ((mword 3 * mword 3)) =>
+ (_s3682_
+ _s3158_) >>= fun w__141 : option ((mword 3 * mword 3)) =>
(match w__141 with
| Some
- ((rsd, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2231_
- _s1697_) >>= fun w__144 : option ((mword 3 * mword 3)) =>
+ (_s3694_
+ _s3158_) >>= fun w__144 : option ((mword 3 * mword 3)) =>
(if
- ((match w__144 with
- | Some
- ((rsd, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__144 with
+ | Some
+ (rsd, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2231_
- _s1697_) >>= fun w__145 : option ((mword 3 * mword 3)) =>
+ (_s3694_
+ _s3158_) >>= fun w__145 : option ((mword 3 * mword 3)) =>
(match w__145 with
| Some
- ((rsd, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2243_
- _s1697_) >>= fun w__148 : option ((mword 3 * mword 3)) =>
+ (_s3706_
+ _s3158_) >>= fun w__148 : option ((mword 3 * mword 3)) =>
(if
- ((match w__148 with
- | Some
- ((rsd, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__148 with
+ | Some
+ (rsd, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s2243_
- _s1697_) >>= fun w__149 : option ((mword 3 * mword 3)) =>
+ (_s3706_
+ _s3158_) >>= fun w__149 : option ((mword 3 * mword 3)) =>
(match w__149 with
| Some
- ((rsd, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2255_
- _s1697_) >>= fun w__152 : option ((mword 3 * mword 3)) =>
+ (_s3718_
+ _s3158_) >>= fun w__152 : option ((mword 3 * mword 3)) =>
(if
- ((match w__152 with
- | Some
- ((rsd, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__152 with
+ | Some
+ (rsd, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s2255_
- _s1697_) >>= fun w__153 : option ((mword 3 * mword 3)) =>
+ (_s3718_
+ _s3158_) >>= fun w__153 : option ((mword 3 * mword 3)) =>
(match w__153 with
| Some
- ((rsd, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2267_
- _s1697_) >>= fun w__156 : option (mword 11) =>
+ (_s3730_
+ _s3158_) >>= fun w__156 : option (mword 11) =>
(if
- ((match w__156 with
- | Some
- (imm) =>
- true
- | _ =>
- false
- end))
+ match w__156 with
+ | Some
+ imm =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2267_
- _s1697_) >>= fun w__157 : option (mword 11) =>
+ (_s3730_
+ _s3158_) >>= fun w__157 : option (mword 11) =>
(match w__157 with
| Some
- (imm) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ imm =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2275_
- _s1697_) >>= fun w__160 : option ((mword 3 * mword 8)) =>
+ (_s3738_
+ _s3158_) >>= fun w__160 : option ((mword 3 * mword 8)) =>
(if
- ((match w__160 with
- | Some
- ((rs, imm)) =>
- true
- | _ =>
- false
- end))
+ match w__160 with
+ | Some
+ (rs, imm) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2275_
- _s1697_) >>= fun w__161 : option ((mword 3 * mword 8)) =>
+ (_s3738_
+ _s3158_) >>= fun w__161 : option ((mword 3 * mword 8)) =>
(match w__161 with
| Some
- ((rs, imm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rs, imm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2287_
- _s1697_) >>= fun w__164 : option ((mword 3 * mword 8)) =>
+ (_s3750_
+ _s3158_) >>= fun w__164 : option ((mword 3 * mword 8)) =>
(if
- ((match w__164 with
- | Some
- ((rs, imm)) =>
- true
- | _ =>
- false
- end))
+ match w__164 with
+ | Some
+ (rs, imm) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2287_
- _s1697_) >>= fun w__165 : option ((mword 3 * mword 8)) =>
+ (_s3750_
+ _s3158_) >>= fun w__165 : option ((mword 3 * mword 8)) =>
(match w__165 with
| Some
- ((rs, imm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rs, imm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2299_
- _s1697_) >>= fun w__168 : option ((mword 5 * mword 6)) =>
+ (_s3762_
+ _s3158_) >>= fun w__168 : option ((mword 5 * mword 6)) =>
(if
- ((match w__168 with
- | Some
- ((rsd, shamt)) =>
- andb
- (neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__168 with
+ | Some
+ (rsd, shamt) =>
+ andb
+ (neq_vec
+ shamt
+ ('b"000000"
+ : mword 6))
+ (neq_vec
+ rsd
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s2299_
- _s1697_) >>= fun w__169 : option ((mword 5 * mword 6)) =>
+ (_s3762_
+ _s3158_) >>= fun w__169 : option ((mword 5 * mword 6)) =>
(match w__169 with
| Some
- ((rsd, shamt)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, shamt) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2311_
- _s1697_) >>= fun w__172 : option ((mword 5 * mword 6)) =>
+ (_s3774_
+ _s3158_) >>= fun w__172 : option ((mword 5 * mword 6)) =>
(if
- ((match w__172 with
- | Some
- ((rd, uimm)) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__172 with
+ | Some
+ (rd, uimm) =>
+ neq_vec
+ rd
+ zreg
+ | _ =>
+ false
+ end
then
- (_s2311_
- _s1697_) >>= fun w__173 : option ((mword 5 * mword 6)) =>
+ (_s3774_
+ _s3158_) >>= fun w__173 : option ((mword 5 * mword 6)) =>
(match w__173 with
| Some
- ((rd, uimm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rd, uimm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2323_
- _s1697_) >>= fun w__176 : option ((mword 5 * mword 6)) =>
+ (_s3786_
+ _s3158_) >>= fun w__176 : option ((mword 5 * mword 6)) =>
(if
- ((match w__176 with
- | Some
- ((rd, uimm)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- (Z.eqb
- 64
- 64)
- | _ =>
- false
- end))
+ match w__176 with
+ | Some
+ (rd, uimm) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
then
- (_s2323_
- _s1697_) >>= fun w__177 : option ((mword 5 * mword 6)) =>
+ (_s3786_
+ _s3158_) >>= fun w__177 : option ((mword 5 * mword 6)) =>
(match w__177 with
| Some
- ((rd, uimm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rd, uimm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2335_
- _s1697_) >>= fun w__180 : option ((mword 5 * mword 6)) =>
+ (_s3798_
+ _s3158_) >>= fun w__180 : option ((mword 5 * mword 6)) =>
(if
- ((match w__180 with
- | Some
- ((rd, uimm)) =>
- true
- | _ =>
- false
- end))
+ match w__180 with
+ | Some
+ (rd, uimm) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2335_
- _s1697_) >>= fun w__181 : option ((mword 5 * mword 6)) =>
+ (_s3798_
+ _s3158_) >>= fun w__181 : option ((mword 5 * mword 6)) =>
(match w__181 with
| Some
- ((rd, uimm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rd, uimm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2347_
- _s1697_) >>= fun w__184 : option ((mword 5 * mword 6)) =>
+ (_s3810_
+ _s3158_) >>= fun w__184 : option ((mword 5 * mword 6)) =>
(if
- ((match w__184 with
- | Some
- ((rs2, uimm)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__184 with
+ | Some
+ (rs2, uimm) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s2347_
- _s1697_) >>= fun w__185 : option ((mword 5 * mword 6)) =>
+ (_s3810_
+ _s3158_) >>= fun w__185 : option ((mword 5 * mword 6)) =>
(match w__185 with
| Some
- ((rs2, uimm)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rs2, uimm) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2359_
- _s1697_) >>= fun w__188 : option (mword 5) =>
+ (_s3822_
+ _s3158_) >>= fun w__188 : option (mword 5) =>
(if
- ((match w__188 with
- | Some
- (rs1) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs1))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__188 with
+ | Some
+ rs1 =>
+ neq_vec
+ rs1
+ zreg
+ | _ =>
+ false
+ end
then
- (_s2359_
- _s1697_) >>= fun w__189 : option (mword 5) =>
+ (_s3822_
+ _s3158_) >>= fun w__189 : option (mword 5) =>
(match w__189 with
| Some
- (rs1) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ rs1 =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2367_
- _s1697_) >>= fun w__192 : option (mword 5) =>
+ (_s3830_
+ _s3158_) >>= fun w__192 : option (mword 5) =>
(if
- ((match w__192 with
- | Some
- (rs1) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs1))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__192 with
+ | Some
+ rs1 =>
+ neq_vec
+ rs1
+ zreg
+ | _ =>
+ false
+ end
then
- (_s2367_
- _s1697_) >>= fun w__193 : option (mword 5) =>
+ (_s3830_
+ _s3158_) >>= fun w__193 : option (mword 5) =>
(match w__193 with
| Some
- (rs1) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ rs1 =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2375_
- _s1697_) >>= fun w__196 : option ((mword 5 * mword 5)) =>
+ (_s3838_
+ _s3158_) >>= fun w__196 : option ((mword 5 * mword 5)) =>
(if
- ((match w__196 with
- | Some
- ((rd, rs2)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs2))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__196 with
+ | Some
+ (rd, rs2) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (neq_vec
+ rs2
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s2375_
- _s1697_) >>= fun w__197 : option ((mword 5 * mword 5)) =>
+ (_s3838_
+ _s3158_) >>= fun w__197 : option ((mword 5 * mword 5)) =>
(match w__197 with
| Some
- ((rd, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rd, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else if
- ((generic_eq
- _s1697_
- "c.ebreak"))
+ generic_eq
+ _s3158_
+ "c.ebreak"
then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ returnm true
else
- (_s2387_
- _s1697_) >>= fun w__200 : option ((mword 5 * mword 5)) =>
+ (_s3850_
+ _s3158_) >>= fun w__200 : option ((mword 5 * mword 5)) =>
(if
- ((match w__200 with
- | Some
- ((rsd, rs2)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs2))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__200 with
+ | Some
+ (rsd, rs2) =>
+ andb
+ (neq_vec
+ rsd
+ zreg)
+ (neq_vec
+ rs2
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s2387_
- _s1697_) >>= fun w__201 : option ((mword 5 * mword 5)) =>
+ (_s3850_
+ _s3158_) >>= fun w__201 : option ((mword 5 * mword 5)) =>
(match w__201 with
| Some
- ((rsd, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rsd, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2399_
- _s1697_) >>= fun w__204 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3862_
+ _s3158_) >>= fun w__204 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__204 with
- | Some
- ((high, signed1, signed2, rd, rs1, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__204 with
+ | Some
+ (high, signed1, signed2, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2399_
- _s1697_) >>= fun w__205 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3862_
+ _s3158_) >>= fun w__205 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5)) =>
(match w__205 with
| Some
- ((high, signed1, signed2, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (high, signed1, signed2, rd, rs1, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2416_
- _s1697_) >>= fun w__208 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3879_
+ _s3158_) >>= fun w__208 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__208 with
- | Some
- ((s, rd, rs1, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__208 with
+ | Some
+ (s, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2416_
- _s1697_) >>= fun w__209 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3879_
+ _s3158_) >>= fun w__209 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(match w__209 with
| Some
- ((s, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (s, rd, rs1, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2434_
- _s1697_) >>= fun w__212 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3897_
+ _s3158_) >>= fun w__212 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__212 with
- | Some
- ((s, rd, rs1, rs2)) =>
- true
- | _ =>
- false
- end))
+ match w__212 with
+ | Some
+ (s, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2434_
- _s1697_) >>= fun w__213 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3897_
+ _s3158_) >>= fun w__213 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(match w__213 with
| Some
- ((s, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (s, rd, rs1, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2452_
- _s1697_) >>= fun w__216 : option ((mword 5 * mword 5 * mword 5)) =>
+ (_s3915_
+ _s3158_) >>= fun w__216 : option ((mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__216 with
- | Some
- ((rd, rs1, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__216 with
+ | Some
+ (rd, rs1, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s2452_
- _s1697_) >>= fun w__217 : option ((mword 5 * mword 5 * mword 5)) =>
+ (_s3915_
+ _s3158_) >>= fun w__217 : option ((mword 5 * mword 5 * mword 5)) =>
(match w__217 with
| Some
- ((rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (rd, rs1, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2468_
- _s1697_) >>= fun w__220 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3931_
+ _s3158_) >>= fun w__220 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__220 with
- | Some
- ((s, rd, rs1, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__220 with
+ | Some
+ (s, rd, rs1, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s2468_
- _s1697_) >>= fun w__221 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3931_
+ _s3158_) >>= fun w__221 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(match w__221 with
| Some
- ((s, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (s, rd, rs1, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2487_
- _s1697_) >>= fun w__224 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3950_
+ _s3158_) >>= fun w__224 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(if
- ((match w__224 with
- | Some
- ((s, rd, rs1, rs2)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__224 with
+ | Some
+ (s, rd, rs1, rs2) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s2487_
- _s1697_) >>= fun w__225 : option ((bool * mword 5 * mword 5 * mword 5)) =>
+ (_s3950_
+ _s3158_) >>= fun w__225 : option ((bool * mword 5 * mword 5 * mword 5)) =>
(match w__225 with
| Some
- ((s, rd, rs1, rs2)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (s, rd, rs1, rs2) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2506_
- _s1697_) >>= fun w__228 : option ((csrop * mword 5 * mword 5 * mword 12)) =>
+ (_s3969_
+ _s3158_) >>= fun w__228 : option ((csrop * mword 5 * mword 12 * mword 5)) =>
(if
- ((match w__228 with
- | Some
- ((op, rd, rs1, csr)) =>
- true
- | _ =>
- false
- end))
+ match w__228 with
+ | Some
+ (op, rd, csr, rs1) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2506_
- _s1697_) >>= fun w__229 : option ((csrop * mword 5 * mword 5 * mword 12)) =>
+ (_s3969_
+ _s3158_) >>= fun w__229 : option ((csrop * mword 5 * mword 12 * mword 5)) =>
(match w__229 with
| Some
- ((op, rd, rs1, csr)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (op, rd, csr, rs1) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2524_
- _s1697_) >>= fun w__232 : option ((csrop * mword 5 * mword 5 * mword 12)) =>
+ (_s3987_
+ _s3158_) >>= fun w__232 : option ((csrop * mword 5 * mword 12 * mword 5)) =>
(if
- ((match w__232 with
- | Some
- ((op, rd, rs1, csr)) =>
- true
- | _ =>
- false
- end))
+ match w__232 with
+ | Some
+ (op, rd, csr, rs1) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2524_
- _s1697_) >>= fun w__233 : option ((csrop * mword 5 * mword 5 * mword 12)) =>
+ (_s3987_
+ _s3158_) >>= fun w__233 : option ((csrop * mword 5 * mword 12 * mword 5)) =>
(match w__233 with
| Some
- ((op, rd, rs1, csr)) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (op, rd, csr, rs1) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else if
- ((generic_eq
- _s1697_
- "uret"))
+ generic_eq
+ _s3158_
+ "uret"
then
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ returnm true
else
- (_s2541_
- _s1697_) >>= fun w__236 : option (mword 32) =>
+ (_s4004_
+ _s3158_) >>= fun w__236 : option ((word_width * mword 5 * mword 12 * mword 5)) =>
(if
- ((match w__236 with
- | Some
- (s) =>
- true
- | _ =>
- false
- end))
+ match w__236 with
+ | Some
+ (width, rd, imm, rs1) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2541_
- _s1697_) >>= fun w__237 : option (mword 32) =>
+ (_s4004_
+ _s3158_) >>= fun w__237 : option ((word_width * mword 5 * mword 12 * mword 5)) =>
(match w__237 with
| Some
- (s) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (width, rd, imm, rs1) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- (_s2549_
- _s1697_) >>= fun w__240 : option (mword 16) =>
+ (_s4028_
+ _s3158_) >>= fun w__240 : option ((word_width * mword 5 * mword 12 * mword 5)) =>
(if
- ((match w__240 with
- | Some
- (s) =>
- true
- | _ =>
- false
- end))
+ match w__240 with
+ | Some
+ (width, rs2, imm, rs1) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s2549_
- _s1697_) >>= fun w__241 : option (mword 16) =>
+ (_s4028_
+ _s3158_) >>= fun w__241 : option ((word_width * mword 5 * mword 12 * mword 5)) =>
(match w__241 with
| Some
- (s) =>
- returnm (projT1
- (build_ex
- true
- : {_bool : bool & ArithFact (_bool =
- true)}))
+ (width, rs2, imm, rs1) =>
+ returnm true
| _ =>
exit tt
: M (bool)
end)
: M (bool)
else
- returnm (projT1
- (build_ex
- false
- : {_bool : bool & ArithFact (not (_bool =
- true))})))
+ (_s4052_
+ _s3158_) >>= fun w__244 : option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__244 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4052_
+ _s3158_) >>= fun w__245 : option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__245 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4077_
+ _s3158_) >>= fun w__248 : option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__248 with
+ | Some
+ (op, rd, rs1, rs2, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4077_
+ _s3158_) >>= fun w__249 : option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__249 with
+ | Some
+ (op, rd, rs1, rs2, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4098_
+ _s3158_) >>= fun w__252 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__252 with
+ | Some
+ (FSQRT_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4098_
+ _s3158_) >>= fun w__253 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__253 with
+ | Some
+ (FSQRT_S, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4115_
+ _s3158_) >>= fun w__256 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__256 with
+ | Some
+ (FCVT_W_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4115_
+ _s3158_) >>= fun w__257 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__257 with
+ | Some
+ (FCVT_W_S, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4132_
+ _s3158_) >>= fun w__260 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__260 with
+ | Some
+ (FCVT_WU_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4132_
+ _s3158_) >>= fun w__261 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__261 with
+ | Some
+ (FCVT_WU_S, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4149_
+ _s3158_) >>= fun w__264 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__264 with
+ | Some
+ (FCVT_S_W, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4149_
+ _s3158_) >>= fun w__265 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__265 with
+ | Some
+ (FCVT_S_W, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4166_
+ _s3158_) >>= fun w__268 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__268 with
+ | Some
+ (FCVT_S_WU, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4166_
+ _s3158_) >>= fun w__269 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__269 with
+ | Some
+ (FCVT_S_WU, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4183_
+ _s3158_) >>= fun w__272 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__272 with
+ | Some
+ (FCVT_L_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4183_
+ _s3158_) >>= fun w__273 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__273 with
+ | Some
+ (FCVT_L_S, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4200_
+ _s3158_) >>= fun w__276 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__276 with
+ | Some
+ (FCVT_LU_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4200_
+ _s3158_) >>= fun w__277 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__277 with
+ | Some
+ (FCVT_LU_S, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4217_
+ _s3158_) >>= fun w__280 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__280 with
+ | Some
+ (FCVT_S_L, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4217_
+ _s3158_) >>= fun w__281 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__281 with
+ | Some
+ (FCVT_S_L, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4234_
+ _s3158_) >>= fun w__284 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__284 with
+ | Some
+ (FCVT_S_LU, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4234_
+ _s3158_) >>= fun w__285 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__285 with
+ | Some
+ (FCVT_S_LU, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4251_
+ _s3158_) >>= fun w__288 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__288 with
+ | Some
+ (FSGNJ_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4251_
+ _s3158_) >>= fun w__289 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__289 with
+ | Some
+ (FSGNJ_S, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4268_
+ _s3158_) >>= fun w__292 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__292 with
+ | Some
+ (FSGNJN_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4268_
+ _s3158_) >>= fun w__293 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__293 with
+ | Some
+ (FSGNJN_S, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4285_
+ _s3158_) >>= fun w__296 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__296 with
+ | Some
+ (FSGNJX_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4285_
+ _s3158_) >>= fun w__297 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__297 with
+ | Some
+ (FSGNJX_S, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4302_
+ _s3158_) >>= fun w__300 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__300 with
+ | Some
+ (FMIN_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4302_
+ _s3158_) >>= fun w__301 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__301 with
+ | Some
+ (FMIN_S, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4319_
+ _s3158_) >>= fun w__304 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__304 with
+ | Some
+ (FMAX_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4319_
+ _s3158_) >>= fun w__305 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__305 with
+ | Some
+ (FMAX_S, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4336_
+ _s3158_) >>= fun w__308 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__308 with
+ | Some
+ (FEQ_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4336_
+ _s3158_) >>= fun w__309 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__309 with
+ | Some
+ (FEQ_S, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4353_
+ _s3158_) >>= fun w__312 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__312 with
+ | Some
+ (FLT_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4353_
+ _s3158_) >>= fun w__313 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__313 with
+ | Some
+ (FLT_S, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4370_
+ _s3158_) >>= fun w__316 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__316 with
+ | Some
+ (FLE_S, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4370_
+ _s3158_) >>= fun w__317 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5)) =>
+ (match w__317 with
+ | Some
+ (FLE_S, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4387_
+ _s3158_) >>= fun w__320 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (if
+ match w__320 with
+ | Some
+ (FMV_X_W, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4387_
+ _s3158_) >>= fun w__321 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (match w__321 with
+ | Some
+ (FMV_X_W, rd, rs1) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4400_
+ _s3158_) >>= fun w__324 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (if
+ match w__324 with
+ | Some
+ (FMV_W_X, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4400_
+ _s3158_) >>= fun w__325 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (match w__325 with
+ | Some
+ (FMV_W_X, rd, rs1) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4413_
+ _s3158_) >>= fun w__328 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (if
+ match w__328 with
+ | Some
+ (FCLASS_S, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4413_
+ _s3158_) >>= fun w__329 : option ((f_un_op_S * mword 5 * mword 5)) =>
+ (match w__329 with
+ | Some
+ (FCLASS_S, rd, rs1) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4426_
+ _s3158_) >>= fun w__332 : option ((mword 5 * mword 6)) =>
+ (if
+ match w__332 with
+ | Some
+ (rd, imm) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s4426_
+ _s3158_) >>= fun w__333 : option ((mword 5 * mword 6)) =>
+ (match w__333 with
+ | Some
+ (rd, imm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4438_
+ _s3158_) >>= fun w__336 : option ((mword 5 * mword 6)) =>
+ (if
+ match w__336 with
+ | Some
+ (rd, uimm) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s4438_
+ _s3158_) >>= fun w__337 : option ((mword 5 * mword 6)) =>
+ (match w__337 with
+ | Some
+ (rd, uimm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4450_
+ _s3158_) >>= fun w__340 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if
+ match w__340 with
+ | Some
+ (rdc, rsc, uimm) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s4450_
+ _s3158_) >>= fun w__341 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__341 with
+ | Some
+ (rdc, rsc, uimm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4466_
+ _s3158_) >>= fun w__344 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if
+ match w__344 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s4466_
+ _s3158_) >>= fun w__345 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__345 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4482_
+ _s3158_) >>= fun w__348 : option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__348 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4482_
+ _s3158_) >>= fun w__349 : option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__349 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4507_
+ _s3158_) >>= fun w__352 : option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__352 with
+ | Some
+ (op, rd, rs1, rs2, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4507_
+ _s3158_) >>= fun w__353 : option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__353 with
+ | Some
+ (op, rd, rs1, rs2, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4528_
+ _s3158_) >>= fun w__356 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__356 with
+ | Some
+ (FSQRT_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4528_
+ _s3158_) >>= fun w__357 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__357 with
+ | Some
+ (FSQRT_D, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4545_
+ _s3158_) >>= fun w__360 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__360 with
+ | Some
+ (FCVT_W_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4545_
+ _s3158_) >>= fun w__361 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__361 with
+ | Some
+ (FCVT_W_D, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4562_
+ _s3158_) >>= fun w__364 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__364 with
+ | Some
+ (FCVT_WU_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4562_
+ _s3158_) >>= fun w__365 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__365 with
+ | Some
+ (FCVT_WU_D, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4579_
+ _s3158_) >>= fun w__368 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__368 with
+ | Some
+ (FCVT_D_W, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4579_
+ _s3158_) >>= fun w__369 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__369 with
+ | Some
+ (FCVT_D_W, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4596_
+ _s3158_) >>= fun w__372 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__372 with
+ | Some
+ (FCVT_D_WU, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4596_
+ _s3158_) >>= fun w__373 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__373 with
+ | Some
+ (FCVT_D_WU, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4613_
+ _s3158_) >>= fun w__376 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__376 with
+ | Some
+ (FCVT_L_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4613_
+ _s3158_) >>= fun w__377 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__377 with
+ | Some
+ (FCVT_L_D, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4630_
+ _s3158_) >>= fun w__380 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__380 with
+ | Some
+ (FCVT_LU_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4630_
+ _s3158_) >>= fun w__381 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__381 with
+ | Some
+ (FCVT_LU_D, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4647_
+ _s3158_) >>= fun w__384 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__384 with
+ | Some
+ (FCVT_D_L, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4647_
+ _s3158_) >>= fun w__385 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__385 with
+ | Some
+ (FCVT_D_L, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4664_
+ _s3158_) >>= fun w__388 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__388 with
+ | Some
+ (FCVT_D_LU, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4664_
+ _s3158_) >>= fun w__389 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__389 with
+ | Some
+ (FCVT_D_LU, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4681_
+ _s3158_) >>= fun w__392 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__392 with
+ | Some
+ (FCVT_S_D, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4681_
+ _s3158_) >>= fun w__393 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__393 with
+ | Some
+ (FCVT_S_D, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4698_
+ _s3158_) >>= fun w__396 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (if
+ match w__396 with
+ | Some
+ (FCVT_D_S, rd, rs1, rm) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4698_
+ _s3158_) >>= fun w__397 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode)) =>
+ (match w__397 with
+ | Some
+ (FCVT_D_S, rd, rs1, rm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4715_
+ _s3158_) >>= fun w__400 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__400 with
+ | Some
+ (FSGNJ_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4715_
+ _s3158_) >>= fun w__401 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__401 with
+ | Some
+ (FSGNJ_D, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4732_
+ _s3158_) >>= fun w__404 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__404 with
+ | Some
+ (FSGNJN_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4732_
+ _s3158_) >>= fun w__405 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__405 with
+ | Some
+ (FSGNJN_D, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4749_
+ _s3158_) >>= fun w__408 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__408 with
+ | Some
+ (FSGNJX_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4749_
+ _s3158_) >>= fun w__409 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__409 with
+ | Some
+ (FSGNJX_D, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4766_
+ _s3158_) >>= fun w__412 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__412 with
+ | Some
+ (FMIN_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4766_
+ _s3158_) >>= fun w__413 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__413 with
+ | Some
+ (FMIN_D, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4783_
+ _s3158_) >>= fun w__416 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__416 with
+ | Some
+ (FMAX_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4783_
+ _s3158_) >>= fun w__417 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__417 with
+ | Some
+ (FMAX_D, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4800_
+ _s3158_) >>= fun w__420 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__420 with
+ | Some
+ (FEQ_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4800_
+ _s3158_) >>= fun w__421 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__421 with
+ | Some
+ (FEQ_D, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4817_
+ _s3158_) >>= fun w__424 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__424 with
+ | Some
+ (FLT_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4817_
+ _s3158_) >>= fun w__425 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__425 with
+ | Some
+ (FLT_D, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4834_
+ _s3158_) >>= fun w__428 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (if
+ match w__428 with
+ | Some
+ (FLE_D, rd, rs1, rs2) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4834_
+ _s3158_) >>= fun w__429 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5)) =>
+ (match w__429 with
+ | Some
+ (FLE_D, rd, rs1, rs2) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4851_
+ _s3158_) >>= fun w__432 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (if
+ match w__432 with
+ | Some
+ (FMV_X_D, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4851_
+ _s3158_) >>= fun w__433 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (match w__433 with
+ | Some
+ (FMV_X_D, rd, rs1) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4864_
+ _s3158_) >>= fun w__436 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (if
+ match w__436 with
+ | Some
+ (FMV_D_X, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4864_
+ _s3158_) >>= fun w__437 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (match w__437 with
+ | Some
+ (FMV_D_X, rd, rs1) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4877_
+ _s3158_) >>= fun w__440 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (if
+ match w__440 with
+ | Some
+ (FCLASS_D, rd, rs1) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4877_
+ _s3158_) >>= fun w__441 : option ((f_un_op_D * mword 5 * mword 5)) =>
+ (match w__441 with
+ | Some
+ (FCLASS_D, rd, rs1) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4890_
+ _s3158_) >>= fun w__444 : option ((mword 5 * mword 6)) =>
+ (if
+ match w__444 with
+ | Some
+ (rd, uimm) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s4890_
+ _s3158_) >>= fun w__445 : option ((mword 5 * mword 6)) =>
+ (match w__445 with
+ | Some
+ (rd, uimm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4902_
+ _s3158_) >>= fun w__448 : option ((mword 5 * mword 6)) =>
+ (if
+ match w__448 with
+ | Some
+ (rs2, uimm) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s4902_
+ _s3158_) >>= fun w__449 : option ((mword 5 * mword 6)) =>
+ (match w__449 with
+ | Some
+ (rs2, uimm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4914_
+ _s3158_) >>= fun w__452 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if
+ match w__452 with
+ | Some
+ (rdc, rsc, uimm) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s4914_
+ _s3158_) >>= fun w__453 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__453 with
+ | Some
+ (rdc, rsc, uimm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4930_
+ _s3158_) >>= fun w__456 : option ((mword 3 * mword 3 * mword 5)) =>
+ (if
+ match w__456 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s4930_
+ _s3158_) >>= fun w__457 : option ((mword 3 * mword 3 * mword 5)) =>
+ (match w__457 with
+ | Some
+ (rsc1, rsc2, uimm) =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4946_
+ _s3158_) >>= fun w__460 : option (mword 32) =>
+ (if
+ match w__460 with
+ | Some
+ s =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4946_
+ _s3158_) >>= fun w__461 : option (mword 32) =>
+ (match w__461 with
+ | Some
+ s =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ (_s4954_
+ _s3158_) >>= fun w__464 : option (mword 16) =>
+ (if
+ match w__464 with
+ | Some
+ s =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s4954_
+ _s3158_) >>= fun w__465 : option (mword 16) =>
+ (match w__465 with
+ | Some
+ s =>
+ returnm true
+ | _ =>
+ exit tt
+ : M (bool)
+ end)
+ : M (bool)
+ else
+ returnm false)
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
+ : M (bool))
: M (bool))
: M (bool))
: M (bool))
@@ -35870,7039 +54397,13127 @@ Definition assembly_backwards_matches (arg_ : string)
: M (bool))
: M (bool).
-Definition _s3457_ (_s3458_ : string)
-: M (option ((mword 16 * string))) :=
-
- let _s3459_ := _s3458_ in
- (if ((string_startswith _s3459_ "c.illegal")) then
- (match (string_drop _s3459_ (projT1 (string_length "c.illegal"))) with
- | _s3460_ =>
- (spc_matches_prefix _s3460_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s3461_ _)) =>
- match (string_drop _s3460_ _s3461_) with
- | _s3462_ =>
- match (hex_bits_16_matches_prefix _s3462_) with
- | Some ((s, existT _ _s3463_ _)) =>
- match (string_drop _s3462_ _s3463_) with | s_ => Some ((s, s_)) end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 16 * string)))
+Definition _s6806_ (_s6807_ : string) : M (option ((mword 16 * string))) :=
+ let _s6808_ := _s6807_ in
+ (if string_startswith _s6808_ "c.illegal" then
+ (match (string_drop _s6808_ (projT1 (string_length "c.illegal"))) with
+ | _s6809_ =>
+ (spc_matches_prefix _s6809_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s6810_ _) =>
+ match (string_drop _s6809_ _s6810_) with
+ | _s6811_ =>
+ match (hex_bits_16_matches_prefix _s6811_) with
+ | Some (s, existT _ _s6812_ _) =>
+ match (string_drop _s6811_ _s6812_) with | s_ => Some (s, s_) end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 16 * string)))
- else returnm (None : option ((mword 16 * string))))
+ else returnm None)
: M (option ((mword 16 * string))).
-Definition _s3449_ (_s3450_ : string)
-: M (option ((mword 32 * string))) :=
-
- let _s3451_ := _s3450_ in
- (if ((string_startswith _s3451_ "illegal")) then
- (match (string_drop _s3451_ (projT1 (string_length "illegal"))) with
- | _s3452_ =>
- (spc_matches_prefix _s3452_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s3453_ _)) =>
- match (string_drop _s3452_ _s3453_) with
- | _s3454_ =>
- match (hex_bits_32_matches_prefix _s3454_) with
- | Some ((s, existT _ _s3455_ _)) =>
- match (string_drop _s3454_ _s3455_) with | s_ => Some ((s, s_)) end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 32 * string)))
+Definition _s6798_ (_s6799_ : string) : M (option ((mword 32 * string))) :=
+ let _s6800_ := _s6799_ in
+ (if string_startswith _s6800_ "illegal" then
+ (match (string_drop _s6800_ (projT1 (string_length "illegal"))) with
+ | _s6801_ =>
+ (spc_matches_prefix _s6801_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s6802_ _) =>
+ match (string_drop _s6801_ _s6802_) with
+ | _s6803_ =>
+ match (hex_bits_32_matches_prefix _s6803_) with
+ | Some (s, existT _ _s6804_ _) =>
+ match (string_drop _s6803_ _s6804_) with | s_ => Some (s, s_) end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 32 * string)))
- else returnm (None : option ((mword 32 * string))))
+ else returnm None)
: M (option ((mword 32 * string))).
-Definition _s3445_ (_s3446_ : string)
-: option string :=
-
- let _s3447_ := _s3446_ in
- if ((string_startswith _s3447_ "uret")) then
- match (string_drop _s3447_ (projT1 (string_length "uret"))) with | s_ => Some (s_) end
- else None.
+Definition _s6782_ (_s6783_ : string) : M (option ((mword 3 * mword 3 * mword 5 * string))) :=
+ let _s6784_ := _s6783_ in
+ (if string_startswith _s6784_ "c.fsd" then
+ (match (string_drop _s6784_ (projT1 (string_length "c.fsd"))) with
+ | _s6785_ =>
+ (spc_matches_prefix _s6785_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s6786_ _) =>
+ (match (string_drop _s6785_ _s6786_) with
+ | _s6787_ =>
+ (creg_name_matches_prefix _s6787_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rsc1, existT _ _s6788_ _) =>
+ (match (string_drop _s6787_ _s6788_) with
+ | _s6789_ =>
+ (sep_matches_prefix _s6789_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s6790_ _) =>
+ (match (string_drop _s6789_ _s6790_) with
+ | _s6791_ =>
+ (creg_name_matches_prefix _s6791_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc2, existT _ _s6792_ _) =>
+ (match (string_drop _s6791_ _s6792_) with
+ | _s6793_ =>
+ (sep_matches_prefix _s6793_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s6794_ _) =>
+ match (string_drop _s6793_ _s6794_) with
+ | _s6795_ =>
+ match (hex_bits_8_matches_prefix _s6795_) with
+ | Some (v__1378, existT _ _s6796_ _) =>
+ if eq_vec (subrange_vec_dec v__1378 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1378 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1378 7 3 in
+ match (string_drop _s6795_ _s6796_) with
+ | s_ => Some (rsc1, rsc2, uimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5 * string))).
+
+Definition _s6766_ (_s6767_ : string) : M (option ((mword 3 * mword 3 * mword 5 * string))) :=
+ let _s6768_ := _s6767_ in
+ (if string_startswith _s6768_ "c.fld" then
+ (match (string_drop _s6768_ (projT1 (string_length "c.fld"))) with
+ | _s6769_ =>
+ (spc_matches_prefix _s6769_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s6770_ _) =>
+ (match (string_drop _s6769_ _s6770_) with
+ | _s6771_ =>
+ (creg_name_matches_prefix _s6771_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rdc, existT _ _s6772_ _) =>
+ (match (string_drop _s6771_ _s6772_) with
+ | _s6773_ =>
+ (sep_matches_prefix _s6773_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s6774_ _) =>
+ (match (string_drop _s6773_ _s6774_) with
+ | _s6775_ =>
+ (creg_name_matches_prefix _s6775_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc, existT _ _s6776_ _) =>
+ (match (string_drop _s6775_ _s6776_) with
+ | _s6777_ =>
+ (sep_matches_prefix _s6777_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s6778_ _) =>
+ match (string_drop _s6777_ _s6778_) with
+ | _s6779_ =>
+ match (hex_bits_8_matches_prefix _s6779_) with
+ | Some (v__1380, existT _ _s6780_ _) =>
+ if eq_vec (subrange_vec_dec v__1380 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1380 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1380 7 3 in
+ match (string_drop _s6779_ _s6780_) with
+ | s_ => Some (rdc, rsc, uimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5 * string))).
+
+Definition _s6754_ (_s6755_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s6756_ := _s6755_ in
+ (if string_startswith _s6756_ "c.fsdsp" then
+ (match (string_drop _s6756_ (projT1 (string_length "c.fsdsp"))) with
+ | _s6757_ =>
+ (spc_matches_prefix _s6757_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s6758_ _) =>
+ (match (string_drop _s6757_ _s6758_) with
+ | _s6759_ =>
+ (reg_name_matches_prefix _s6759_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rs2, existT _ _s6760_ _) =>
+ (match (string_drop _s6759_ _s6760_) with
+ | _s6761_ =>
+ (sep_matches_prefix _s6761_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s6762_ _) =>
+ match (string_drop _s6761_ _s6762_) with
+ | _s6763_ =>
+ match (hex_bits_6_matches_prefix _s6763_) with
+ | Some (uimm, existT _ _s6764_ _) =>
+ match (string_drop _s6763_ _s6764_) with
+ | s_ => Some (rs2, uimm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6 * string))).
+
+Definition _s6742_ (_s6743_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s6744_ := _s6743_ in
+ (if string_startswith _s6744_ "c.fldsp" then
+ (match (string_drop _s6744_ (projT1 (string_length "c.fldsp"))) with
+ | _s6745_ =>
+ (spc_matches_prefix _s6745_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s6746_ _) =>
+ (match (string_drop _s6745_ _s6746_) with
+ | _s6747_ =>
+ (reg_name_matches_prefix _s6747_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s6748_ _) =>
+ (match (string_drop _s6747_ _s6748_) with
+ | _s6749_ =>
+ (sep_matches_prefix _s6749_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s6750_ _) =>
+ match (string_drop _s6749_ _s6750_) with
+ | _s6751_ =>
+ match (hex_bits_6_matches_prefix _s6751_) with
+ | Some (uimm, existT _ _s6752_ _) =>
+ match (string_drop _s6751_ _s6752_) with
+ | s_ => Some (rd, uimm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6 * string))).
+
+Definition _s6729_ (_s6730_ : string) : M (option ((f_un_op_D * mword 5 * mword 5 * string))) :=
+ (match _s6730_ with
+ | _s6731_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s6731_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCLASS_D, existT _ _s6732_ _) =>
+ (match (string_drop _s6731_ _s6732_) with
+ | _s6733_ =>
+ (spc_matches_prefix _s6733_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6734_ _) =>
+ (match (string_drop _s6733_ _s6734_) with
+ | _s6735_ =>
+ (reg_name_matches_prefix _s6735_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6736_ _) =>
+ (match (string_drop _s6735_ _s6736_) with
+ | _s6737_ =>
+ (sep_matches_prefix _s6737_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6738_ _) =>
+ (match (string_drop _s6737_ _s6738_) with
+ | _s6739_ =>
+ (freg_name_matches_prefix _s6739_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s6740_ _) =>
+ match (string_drop _s6739_ _s6740_) with
+ | s_ => Some (FCLASS_D, rd, rs1, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string))).
+
+Definition _s6716_ (_s6717_ : string) : M (option ((f_un_op_D * mword 5 * mword 5 * string))) :=
+ (match _s6717_ with
+ | _s6718_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s6718_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_D_X, existT _ _s6719_ _) =>
+ (match (string_drop _s6718_ _s6719_) with
+ | _s6720_ =>
+ (spc_matches_prefix _s6720_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6721_ _) =>
+ (match (string_drop _s6720_ _s6721_) with
+ | _s6722_ =>
+ (freg_name_matches_prefix _s6722_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6723_ _) =>
+ (match (string_drop _s6722_ _s6723_) with
+ | _s6724_ =>
+ (sep_matches_prefix _s6724_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6725_ _) =>
+ (match (string_drop _s6724_ _s6725_) with
+ | _s6726_ =>
+ (reg_name_matches_prefix _s6726_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s6727_ _) =>
+ match (string_drop _s6726_ _s6727_) with
+ | s_ => Some (FMV_D_X, rd, rs1, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string))).
+
+Definition _s6703_ (_s6704_ : string) : M (option ((f_un_op_D * mword 5 * mword 5 * string))) :=
+ (match _s6704_ with
+ | _s6705_ =>
+ (f_un_type_mnemonic_D_matches_prefix _s6705_) >>= fun w__0 : option ((f_un_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_X_D, existT _ _s6706_ _) =>
+ (match (string_drop _s6705_ _s6706_) with
+ | _s6707_ =>
+ (spc_matches_prefix _s6707_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6708_ _) =>
+ (match (string_drop _s6707_ _s6708_) with
+ | _s6709_ =>
+ (reg_name_matches_prefix _s6709_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6710_ _) =>
+ (match (string_drop _s6709_ _s6710_) with
+ | _s6711_ =>
+ (sep_matches_prefix _s6711_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6712_ _) =>
+ (match (string_drop _s6711_ _s6712_) with
+ | _s6713_ =>
+ (freg_name_matches_prefix _s6713_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s6714_ _) =>
+ match (string_drop _s6713_ _s6714_) with
+ | s_ => Some (FMV_X_D, rd, rs1, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_D * mword 5 * mword 5 * string))).
-Definition _s3428_ (_s3429_ : string)
-: M (option ((csrop * mword 5 * mword 5 * mword 12 * string))) :=
-
- (match _s3429_ with
- | _s3430_ =>
- (csr_mnemonic_matches_prefix _s3430_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=
+Definition _s6686_ (_s6687_ : string)
+: M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6687_ with
+ | _s6688_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s6688_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s3431_ _)) =>
- (match (string_drop _s3430_ _s3431_) with
- | _s3432_ =>
- (spc_matches_prefix _s3432_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (FLE_D, existT _ _s6689_ _) =>
+ (match (string_drop _s6688_ _s6689_) with
+ | _s6690_ =>
+ (spc_matches_prefix _s6690_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s3433_ _)) =>
- (match (string_drop _s3432_ _s3433_) with
- | _s3434_ =>
- (reg_name_matches_prefix _s3434_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s6691_ _) =>
+ (match (string_drop _s6690_ _s6691_) with
+ | _s6692_ =>
+ (reg_name_matches_prefix _s6692_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s3435_ _)) =>
- (match (string_drop _s3434_ _s3435_) with
- | _s3436_ =>
- (sep_matches_prefix _s3436_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s6693_ _) =>
+ (match (string_drop _s6692_ _s6693_) with
+ | _s6694_ =>
+ (sep_matches_prefix _s6694_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s3437_ _)) =>
- (match (string_drop _s3436_ _s3437_) with
- | _s3438_ =>
- (reg_name_matches_prefix _s3438_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s6695_ _) =>
+ (match (string_drop _s6694_ _s6695_) with
+ | _s6696_ =>
+ (freg_name_matches_prefix _s6696_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s3439_ _)) =>
- (match (string_drop _s3438_ _s3439_) with
- | _s3440_ =>
- (sep_matches_prefix _s3440_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s6697_ _) =>
+ (match (string_drop _s6696_ _s6697_) with
+ | _s6698_ =>
+ (sep_matches_prefix _s6698_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s3441_ _)) =>
- (match (string_drop _s3440_ _s3441_) with
- | _s3442_ =>
- (csr_name_map_matches_prefix _s3442_) >>= fun w__6 : option ((mword 12 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s6699_ _) =>
+ (match (string_drop _s6698_ _s6699_) with
+ | _s6700_ =>
+ (freg_name_matches_prefix _s6700_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((csr, existT _ _s3443_ _)) =>
- match (string_drop _s3442_
- _s3443_) with
- | s_ =>
- Some ((op, rd, rs1, csr, s_))
- end
- | _ => None
- end)
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6701_ _) =>
+ match (string_drop _s6700_
+ _s6701_) with
+ | s_ =>
+ Some (FLE_D, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6669_ (_s6670_ : string)
+: M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6670_ with
+ | _s6671_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s6671_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLT_D, existT _ _s6672_ _) =>
+ (match (string_drop _s6671_ _s6672_) with
+ | _s6673_ =>
+ (spc_matches_prefix _s6673_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6674_ _) =>
+ (match (string_drop _s6673_ _s6674_) with
+ | _s6675_ =>
+ (reg_name_matches_prefix _s6675_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6676_ _) =>
+ (match (string_drop _s6675_ _s6676_) with
+ | _s6677_ =>
+ (sep_matches_prefix _s6677_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6678_ _) =>
+ (match (string_drop _s6677_ _s6678_) with
+ | _s6679_ =>
+ (freg_name_matches_prefix _s6679_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6680_ _) =>
+ (match (string_drop _s6679_ _s6680_) with
+ | _s6681_ =>
+ (sep_matches_prefix _s6681_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6682_ _) =>
+ (match (string_drop _s6681_ _s6682_) with
+ | _s6683_ =>
+ (freg_name_matches_prefix _s6683_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6684_ _) =>
+ match (string_drop _s6683_
+ _s6684_) with
+ | s_ =>
+ Some (FLT_D, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6652_ (_s6653_ : string)
+: M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6653_ with
+ | _s6654_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s6654_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FEQ_D, existT _ _s6655_ _) =>
+ (match (string_drop _s6654_ _s6655_) with
+ | _s6656_ =>
+ (spc_matches_prefix _s6656_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6657_ _) =>
+ (match (string_drop _s6656_ _s6657_) with
+ | _s6658_ =>
+ (reg_name_matches_prefix _s6658_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6659_ _) =>
+ (match (string_drop _s6658_ _s6659_) with
+ | _s6660_ =>
+ (sep_matches_prefix _s6660_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6661_ _) =>
+ (match (string_drop _s6660_ _s6661_) with
+ | _s6662_ =>
+ (freg_name_matches_prefix _s6662_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6663_ _) =>
+ (match (string_drop _s6662_ _s6663_) with
+ | _s6664_ =>
+ (sep_matches_prefix _s6664_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6665_ _) =>
+ (match (string_drop _s6664_ _s6665_) with
+ | _s6666_ =>
+ (freg_name_matches_prefix _s6666_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6667_ _) =>
+ match (string_drop _s6666_
+ _s6667_) with
+ | s_ =>
+ Some (FEQ_D, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6635_ (_s6636_ : string)
+: M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6636_ with
+ | _s6637_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s6637_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMAX_D, existT _ _s6638_ _) =>
+ (match (string_drop _s6637_ _s6638_) with
+ | _s6639_ =>
+ (spc_matches_prefix _s6639_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6640_ _) =>
+ (match (string_drop _s6639_ _s6640_) with
+ | _s6641_ =>
+ (freg_name_matches_prefix _s6641_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6642_ _) =>
+ (match (string_drop _s6641_ _s6642_) with
+ | _s6643_ =>
+ (sep_matches_prefix _s6643_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6644_ _) =>
+ (match (string_drop _s6643_ _s6644_) with
+ | _s6645_ =>
+ (freg_name_matches_prefix _s6645_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6646_ _) =>
+ (match (string_drop _s6645_ _s6646_) with
+ | _s6647_ =>
+ (sep_matches_prefix _s6647_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6648_ _) =>
+ (match (string_drop _s6647_ _s6648_) with
+ | _s6649_ =>
+ (freg_name_matches_prefix _s6649_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6650_ _) =>
+ match (string_drop _s6649_
+ _s6650_) with
+ | s_ =>
+ Some
+ (FMAX_D, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6618_ (_s6619_ : string)
+: M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6619_ with
+ | _s6620_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s6620_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMIN_D, existT _ _s6621_ _) =>
+ (match (string_drop _s6620_ _s6621_) with
+ | _s6622_ =>
+ (spc_matches_prefix _s6622_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6623_ _) =>
+ (match (string_drop _s6622_ _s6623_) with
+ | _s6624_ =>
+ (freg_name_matches_prefix _s6624_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6625_ _) =>
+ (match (string_drop _s6624_ _s6625_) with
+ | _s6626_ =>
+ (sep_matches_prefix _s6626_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6627_ _) =>
+ (match (string_drop _s6626_ _s6627_) with
+ | _s6628_ =>
+ (freg_name_matches_prefix _s6628_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6629_ _) =>
+ (match (string_drop _s6628_ _s6629_) with
+ | _s6630_ =>
+ (sep_matches_prefix _s6630_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6631_ _) =>
+ (match (string_drop _s6630_ _s6631_) with
+ | _s6632_ =>
+ (freg_name_matches_prefix _s6632_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6633_ _) =>
+ match (string_drop _s6632_
+ _s6633_) with
+ | s_ =>
+ Some
+ (FMIN_D, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6601_ (_s6602_ : string)
+: M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6602_ with
+ | _s6603_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s6603_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJX_D, existT _ _s6604_ _) =>
+ (match (string_drop _s6603_ _s6604_) with
+ | _s6605_ =>
+ (spc_matches_prefix _s6605_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6606_ _) =>
+ (match (string_drop _s6605_ _s6606_) with
+ | _s6607_ =>
+ (freg_name_matches_prefix _s6607_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6608_ _) =>
+ (match (string_drop _s6607_ _s6608_) with
+ | _s6609_ =>
+ (sep_matches_prefix _s6609_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6610_ _) =>
+ (match (string_drop _s6609_ _s6610_) with
+ | _s6611_ =>
+ (freg_name_matches_prefix _s6611_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6612_ _) =>
+ (match (string_drop _s6611_ _s6612_) with
+ | _s6613_ =>
+ (sep_matches_prefix _s6613_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6614_ _) =>
+ (match (string_drop _s6613_ _s6614_) with
+ | _s6615_ =>
+ (freg_name_matches_prefix _s6615_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6616_ _) =>
+ match (string_drop _s6615_
+ _s6616_) with
+ | s_ =>
+ Some
+ (FSGNJX_D, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6584_ (_s6585_ : string)
+: M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6585_ with
+ | _s6586_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s6586_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJN_D, existT _ _s6587_ _) =>
+ (match (string_drop _s6586_ _s6587_) with
+ | _s6588_ =>
+ (spc_matches_prefix _s6588_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6589_ _) =>
+ (match (string_drop _s6588_ _s6589_) with
+ | _s6590_ =>
+ (freg_name_matches_prefix _s6590_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6591_ _) =>
+ (match (string_drop _s6590_ _s6591_) with
+ | _s6592_ =>
+ (sep_matches_prefix _s6592_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6593_ _) =>
+ (match (string_drop _s6592_ _s6593_) with
+ | _s6594_ =>
+ (freg_name_matches_prefix _s6594_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6595_ _) =>
+ (match (string_drop _s6594_ _s6595_) with
+ | _s6596_ =>
+ (sep_matches_prefix _s6596_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6597_ _) =>
+ (match (string_drop _s6596_ _s6597_) with
+ | _s6598_ =>
+ (freg_name_matches_prefix _s6598_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6599_ _) =>
+ match (string_drop _s6598_
+ _s6599_) with
+ | s_ =>
+ Some
+ (FSGNJN_D, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6567_ (_s6568_ : string)
+: M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6568_ with
+ | _s6569_ =>
+ (f_bin_type_mnemonic_D_matches_prefix _s6569_) >>= fun w__0 : option ((f_bin_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJ_D, existT _ _s6570_ _) =>
+ (match (string_drop _s6569_ _s6570_) with
+ | _s6571_ =>
+ (spc_matches_prefix _s6571_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6572_ _) =>
+ (match (string_drop _s6571_ _s6572_) with
+ | _s6573_ =>
+ (freg_name_matches_prefix _s6573_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6574_ _) =>
+ (match (string_drop _s6573_ _s6574_) with
+ | _s6575_ =>
+ (sep_matches_prefix _s6575_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6576_ _) =>
+ (match (string_drop _s6575_ _s6576_) with
+ | _s6577_ =>
+ (freg_name_matches_prefix _s6577_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6578_ _) =>
+ (match (string_drop _s6577_ _s6578_) with
+ | _s6579_ =>
+ (sep_matches_prefix _s6579_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6580_ _) =>
+ (match (string_drop _s6579_ _s6580_) with
+ | _s6581_ =>
+ (freg_name_matches_prefix _s6581_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6582_ _) =>
+ match (string_drop _s6581_
+ _s6582_) with
+ | s_ =>
+ Some
+ (FSGNJ_D, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6550_ (_s6551_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6551_ with
+ | _s6552_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6552_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_S, existT _ _s6553_ _) =>
+ (match (string_drop _s6552_ _s6553_) with
+ | _s6554_ =>
+ (spc_matches_prefix _s6554_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6555_ _) =>
+ (match (string_drop _s6554_ _s6555_) with
+ | _s6556_ =>
+ (freg_name_matches_prefix _s6556_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6557_ _) =>
+ (match (string_drop _s6556_ _s6557_) with
+ | _s6558_ =>
+ (sep_matches_prefix _s6558_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6559_ _) =>
+ (match (string_drop _s6558_ _s6559_) with
+ | _s6560_ =>
+ (freg_name_matches_prefix _s6560_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6561_ _) =>
+ (match (string_drop _s6560_ _s6561_) with
+ | _s6562_ =>
+ (sep_matches_prefix _s6562_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6563_ _) =>
+ (match (string_drop _s6562_ _s6563_) with
+ | _s6564_ =>
+ (frm_mnemonic_matches_prefix _s6564_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6565_ _) =>
+ match (string_drop _s6564_
+ _s6565_) with
+ | s_ =>
+ Some
+ (FCVT_D_S, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6533_ (_s6534_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6534_ with
+ | _s6535_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6535_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_D, existT _ _s6536_ _) =>
+ (match (string_drop _s6535_ _s6536_) with
+ | _s6537_ =>
+ (spc_matches_prefix _s6537_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6538_ _) =>
+ (match (string_drop _s6537_ _s6538_) with
+ | _s6539_ =>
+ (freg_name_matches_prefix _s6539_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6540_ _) =>
+ (match (string_drop _s6539_ _s6540_) with
+ | _s6541_ =>
+ (sep_matches_prefix _s6541_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6542_ _) =>
+ (match (string_drop _s6541_ _s6542_) with
+ | _s6543_ =>
+ (freg_name_matches_prefix _s6543_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6544_ _) =>
+ (match (string_drop _s6543_ _s6544_) with
+ | _s6545_ =>
+ (sep_matches_prefix _s6545_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6546_ _) =>
+ (match (string_drop _s6545_ _s6546_) with
+ | _s6547_ =>
+ (frm_mnemonic_matches_prefix _s6547_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6548_ _) =>
+ match (string_drop _s6547_
+ _s6548_) with
+ | s_ =>
+ Some
+ (FCVT_S_D, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6516_ (_s6517_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6517_ with
+ | _s6518_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6518_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_LU, existT _ _s6519_ _) =>
+ (match (string_drop _s6518_ _s6519_) with
+ | _s6520_ =>
+ (spc_matches_prefix _s6520_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6521_ _) =>
+ (match (string_drop _s6520_ _s6521_) with
+ | _s6522_ =>
+ (freg_name_matches_prefix _s6522_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6523_ _) =>
+ (match (string_drop _s6522_ _s6523_) with
+ | _s6524_ =>
+ (sep_matches_prefix _s6524_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6525_ _) =>
+ (match (string_drop _s6524_ _s6525_) with
+ | _s6526_ =>
+ (reg_name_matches_prefix _s6526_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6527_ _) =>
+ (match (string_drop _s6526_ _s6527_) with
+ | _s6528_ =>
+ (sep_matches_prefix _s6528_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6529_ _) =>
+ (match (string_drop _s6528_ _s6529_) with
+ | _s6530_ =>
+ (frm_mnemonic_matches_prefix _s6530_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6531_ _) =>
+ match (string_drop _s6530_
+ _s6531_) with
+ | s_ =>
+ Some
+ (FCVT_D_LU, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6499_ (_s6500_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6500_ with
+ | _s6501_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6501_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_L, existT _ _s6502_ _) =>
+ (match (string_drop _s6501_ _s6502_) with
+ | _s6503_ =>
+ (spc_matches_prefix _s6503_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6504_ _) =>
+ (match (string_drop _s6503_ _s6504_) with
+ | _s6505_ =>
+ (freg_name_matches_prefix _s6505_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6506_ _) =>
+ (match (string_drop _s6505_ _s6506_) with
+ | _s6507_ =>
+ (sep_matches_prefix _s6507_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6508_ _) =>
+ (match (string_drop _s6507_ _s6508_) with
+ | _s6509_ =>
+ (reg_name_matches_prefix _s6509_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6510_ _) =>
+ (match (string_drop _s6509_ _s6510_) with
+ | _s6511_ =>
+ (sep_matches_prefix _s6511_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6512_ _) =>
+ (match (string_drop _s6511_ _s6512_) with
+ | _s6513_ =>
+ (frm_mnemonic_matches_prefix _s6513_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6514_ _) =>
+ match (string_drop _s6513_
+ _s6514_) with
+ | s_ =>
+ Some
+ (FCVT_D_L, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6482_ (_s6483_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6483_ with
+ | _s6484_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6484_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_LU_D, existT _ _s6485_ _) =>
+ (match (string_drop _s6484_ _s6485_) with
+ | _s6486_ =>
+ (spc_matches_prefix _s6486_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6487_ _) =>
+ (match (string_drop _s6486_ _s6487_) with
+ | _s6488_ =>
+ (reg_name_matches_prefix _s6488_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6489_ _) =>
+ (match (string_drop _s6488_ _s6489_) with
+ | _s6490_ =>
+ (sep_matches_prefix _s6490_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6491_ _) =>
+ (match (string_drop _s6490_ _s6491_) with
+ | _s6492_ =>
+ (freg_name_matches_prefix _s6492_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6493_ _) =>
+ (match (string_drop _s6492_ _s6493_) with
+ | _s6494_ =>
+ (sep_matches_prefix _s6494_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6495_ _) =>
+ (match (string_drop _s6494_ _s6495_) with
+ | _s6496_ =>
+ (frm_mnemonic_matches_prefix _s6496_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6497_ _) =>
+ match (string_drop _s6496_
+ _s6497_) with
+ | s_ =>
+ Some
+ (FCVT_LU_D, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6465_ (_s6466_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6466_ with
+ | _s6467_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6467_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_L_D, existT _ _s6468_ _) =>
+ (match (string_drop _s6467_ _s6468_) with
+ | _s6469_ =>
+ (spc_matches_prefix _s6469_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6470_ _) =>
+ (match (string_drop _s6469_ _s6470_) with
+ | _s6471_ =>
+ (reg_name_matches_prefix _s6471_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6472_ _) =>
+ (match (string_drop _s6471_ _s6472_) with
+ | _s6473_ =>
+ (sep_matches_prefix _s6473_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6474_ _) =>
+ (match (string_drop _s6473_ _s6474_) with
+ | _s6475_ =>
+ (freg_name_matches_prefix _s6475_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6476_ _) =>
+ (match (string_drop _s6475_ _s6476_) with
+ | _s6477_ =>
+ (sep_matches_prefix _s6477_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6478_ _) =>
+ (match (string_drop _s6477_ _s6478_) with
+ | _s6479_ =>
+ (frm_mnemonic_matches_prefix _s6479_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6480_ _) =>
+ match (string_drop _s6479_
+ _s6480_) with
+ | s_ =>
+ Some
+ (FCVT_L_D, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6448_ (_s6449_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6449_ with
+ | _s6450_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6450_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_WU, existT _ _s6451_ _) =>
+ (match (string_drop _s6450_ _s6451_) with
+ | _s6452_ =>
+ (spc_matches_prefix _s6452_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6453_ _) =>
+ (match (string_drop _s6452_ _s6453_) with
+ | _s6454_ =>
+ (freg_name_matches_prefix _s6454_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6455_ _) =>
+ (match (string_drop _s6454_ _s6455_) with
+ | _s6456_ =>
+ (sep_matches_prefix _s6456_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6457_ _) =>
+ (match (string_drop _s6456_ _s6457_) with
+ | _s6458_ =>
+ (reg_name_matches_prefix _s6458_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6459_ _) =>
+ (match (string_drop _s6458_ _s6459_) with
+ | _s6460_ =>
+ (sep_matches_prefix _s6460_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6461_ _) =>
+ (match (string_drop _s6460_ _s6461_) with
+ | _s6462_ =>
+ (frm_mnemonic_matches_prefix _s6462_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6463_ _) =>
+ match (string_drop _s6462_
+ _s6463_) with
+ | s_ =>
+ Some
+ (FCVT_D_WU, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6431_ (_s6432_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6432_ with
+ | _s6433_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6433_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_D_W, existT _ _s6434_ _) =>
+ (match (string_drop _s6433_ _s6434_) with
+ | _s6435_ =>
+ (spc_matches_prefix _s6435_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6436_ _) =>
+ (match (string_drop _s6435_ _s6436_) with
+ | _s6437_ =>
+ (freg_name_matches_prefix _s6437_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6438_ _) =>
+ (match (string_drop _s6437_ _s6438_) with
+ | _s6439_ =>
+ (sep_matches_prefix _s6439_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6440_ _) =>
+ (match (string_drop _s6439_ _s6440_) with
+ | _s6441_ =>
+ (reg_name_matches_prefix _s6441_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6442_ _) =>
+ (match (string_drop _s6441_ _s6442_) with
+ | _s6443_ =>
+ (sep_matches_prefix _s6443_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6444_ _) =>
+ (match (string_drop _s6443_ _s6444_) with
+ | _s6445_ =>
+ (frm_mnemonic_matches_prefix _s6445_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6446_ _) =>
+ match (string_drop _s6445_
+ _s6446_) with
+ | s_ =>
+ Some
+ (FCVT_D_W, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6414_ (_s6415_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6415_ with
+ | _s6416_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6416_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_WU_D, existT _ _s6417_ _) =>
+ (match (string_drop _s6416_ _s6417_) with
+ | _s6418_ =>
+ (spc_matches_prefix _s6418_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6419_ _) =>
+ (match (string_drop _s6418_ _s6419_) with
+ | _s6420_ =>
+ (reg_name_matches_prefix _s6420_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6421_ _) =>
+ (match (string_drop _s6420_ _s6421_) with
+ | _s6422_ =>
+ (sep_matches_prefix _s6422_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6423_ _) =>
+ (match (string_drop _s6422_ _s6423_) with
+ | _s6424_ =>
+ (freg_name_matches_prefix _s6424_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6425_ _) =>
+ (match (string_drop _s6424_ _s6425_) with
+ | _s6426_ =>
+ (sep_matches_prefix _s6426_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6427_ _) =>
+ (match (string_drop _s6426_ _s6427_) with
+ | _s6428_ =>
+ (frm_mnemonic_matches_prefix _s6428_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6429_ _) =>
+ match (string_drop _s6428_
+ _s6429_) with
+ | s_ =>
+ Some
+ (FCVT_WU_D, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6397_ (_s6398_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6398_ with
+ | _s6399_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6399_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_W_D, existT _ _s6400_ _) =>
+ (match (string_drop _s6399_ _s6400_) with
+ | _s6401_ =>
+ (spc_matches_prefix _s6401_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6402_ _) =>
+ (match (string_drop _s6401_ _s6402_) with
+ | _s6403_ =>
+ (reg_name_matches_prefix _s6403_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6404_ _) =>
+ (match (string_drop _s6403_ _s6404_) with
+ | _s6405_ =>
+ (sep_matches_prefix _s6405_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6406_ _) =>
+ (match (string_drop _s6405_ _s6406_) with
+ | _s6407_ =>
+ (freg_name_matches_prefix _s6407_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6408_ _) =>
+ (match (string_drop _s6407_ _s6408_) with
+ | _s6409_ =>
+ (sep_matches_prefix _s6409_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6410_ _) =>
+ (match (string_drop _s6409_ _s6410_) with
+ | _s6411_ =>
+ (frm_mnemonic_matches_prefix _s6411_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6412_ _) =>
+ match (string_drop _s6411_
+ _s6412_) with
+ | s_ =>
+ Some
+ (FCVT_W_D, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6380_ (_s6381_ : string)
+: M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6381_ with
+ | _s6382_ =>
+ (f_un_rm_type_mnemonic_D_matches_prefix _s6382_) >>= fun w__0 : option ((f_un_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSQRT_D, existT _ _s6383_ _) =>
+ (match (string_drop _s6382_ _s6383_) with
+ | _s6384_ =>
+ (spc_matches_prefix _s6384_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6385_ _) =>
+ (match (string_drop _s6384_ _s6385_) with
+ | _s6386_ =>
+ (freg_name_matches_prefix _s6386_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6387_ _) =>
+ (match (string_drop _s6386_ _s6387_) with
+ | _s6388_ =>
+ (sep_matches_prefix _s6388_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6389_ _) =>
+ (match (string_drop _s6388_ _s6389_) with
+ | _s6390_ =>
+ (freg_name_matches_prefix _s6390_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6391_ _) =>
+ (match (string_drop _s6390_ _s6391_) with
+ | _s6392_ =>
+ (sep_matches_prefix _s6392_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6393_ _) =>
+ (match (string_drop _s6392_ _s6393_) with
+ | _s6394_ =>
+ (frm_mnemonic_matches_prefix _s6394_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6395_ _) =>
+ match (string_drop _s6394_
+ _s6395_) with
+ | s_ =>
+ Some
+ (FSQRT_D, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6359_ (_s6360_ : string)
+: M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6360_ with
+ | _s6361_ =>
+ (f_bin_rm_type_mnemonic_D_matches_prefix _s6361_) >>= fun w__0 : option ((f_bin_rm_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s6362_ _) =>
+ (match (string_drop _s6361_ _s6362_) with
+ | _s6363_ =>
+ (spc_matches_prefix _s6363_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6364_ _) =>
+ (match (string_drop _s6363_ _s6364_) with
+ | _s6365_ =>
+ (freg_name_matches_prefix _s6365_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6366_ _) =>
+ (match (string_drop _s6365_ _s6366_) with
+ | _s6367_ =>
+ (sep_matches_prefix _s6367_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6368_ _) =>
+ (match (string_drop _s6367_ _s6368_) with
+ | _s6369_ =>
+ (freg_name_matches_prefix _s6369_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6370_ _) =>
+ (match (string_drop _s6369_ _s6370_) with
+ | _s6371_ =>
+ (sep_matches_prefix _s6371_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6372_ _) =>
+ (match (string_drop _s6371_ _s6372_) with
+ | _s6373_ =>
+ (freg_name_matches_prefix _s6373_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s6374_ _) =>
+ (match (string_drop _s6373_ _s6374_) with
+ | _s6375_ =>
+ (sep_matches_prefix _s6375_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s6376_ _) =>
+ (match (string_drop _s6375_
+ _s6376_) with
+ | _s6377_ =>
+ (frm_mnemonic_matches_prefix
+ _s6377_) >>= fun w__8 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__8 with
+ | Some
+ (rm, existT _ _s6378_ _) =>
+ match (string_drop
+ _s6377_
+ _s6378_) with
+ | s_ =>
+ Some
+ (op, rd, rs1, rs2, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6334_ (_s6335_ : string)
+: M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6335_ with
+ | _s6336_ =>
+ (f_madd_type_mnemonic_D_matches_prefix _s6336_) >>= fun w__0 : option ((f_madd_op_D * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s6337_ _) =>
+ (match (string_drop _s6336_ _s6337_) with
+ | _s6338_ =>
+ (spc_matches_prefix _s6338_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6339_ _) =>
+ (match (string_drop _s6338_ _s6339_) with
+ | _s6340_ =>
+ (freg_name_matches_prefix _s6340_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6341_ _) =>
+ (match (string_drop _s6340_ _s6341_) with
+ | _s6342_ =>
+ (sep_matches_prefix _s6342_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6343_ _) =>
+ (match (string_drop _s6342_ _s6343_) with
+ | _s6344_ =>
+ (freg_name_matches_prefix _s6344_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6345_ _) =>
+ (match (string_drop _s6344_ _s6345_) with
+ | _s6346_ =>
+ (sep_matches_prefix _s6346_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6347_ _) =>
+ (match (string_drop _s6346_ _s6347_) with
+ | _s6348_ =>
+ (freg_name_matches_prefix _s6348_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s6349_ _) =>
+ (match (string_drop _s6348_ _s6349_) with
+ | _s6350_ =>
+ (sep_matches_prefix _s6350_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s6351_ _) =>
+ (match (string_drop _s6350_
+ _s6351_) with
+ | _s6352_ =>
+ (freg_name_matches_prefix
+ _s6352_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__8 with
+ | Some
+ (rs3, existT _ _s6353_ _) =>
+ (match (string_drop
+ _s6352_ _s6353_) with
+ | _s6354_ =>
+ (sep_matches_prefix
+ _s6354_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__9 with
+ | Some
+ (tt, existT _ _s6355_ _) =>
+ (match (string_drop
+ _s6354_
+ _s6355_) with
+ | _s6356_ =>
+ (frm_mnemonic_matches_prefix
+ _s6356_) >>= fun w__10 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__10 with
+ | Some
+ (rm, existT _ _s6357_ _) =>
+ match (string_drop
+ _s6356_
+ _s6357_) with
+ | s_ =>
+ Some
+ (op, rd, rs1, rs2, rs3, rm, s_)
+ end
+ | _ =>
+ None
+ end)
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6318_ (_s6319_ : string) : M (option ((mword 3 * mword 3 * mword 5 * string))) :=
+ let _s6320_ := _s6319_ in
+ (if string_startswith _s6320_ "c.fsw" then
+ (match (string_drop _s6320_ (projT1 (string_length "c.fsw"))) with
+ | _s6321_ =>
+ (spc_matches_prefix _s6321_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s6322_ _) =>
+ (match (string_drop _s6321_ _s6322_) with
+ | _s6323_ =>
+ (creg_name_matches_prefix _s6323_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rsc1, existT _ _s6324_ _) =>
+ (match (string_drop _s6323_ _s6324_) with
+ | _s6325_ =>
+ (sep_matches_prefix _s6325_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s6326_ _) =>
+ (match (string_drop _s6325_ _s6326_) with
+ | _s6327_ =>
+ (creg_name_matches_prefix _s6327_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc2, existT _ _s6328_ _) =>
+ (match (string_drop _s6327_ _s6328_) with
+ | _s6329_ =>
+ (sep_matches_prefix _s6329_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s6330_ _) =>
+ match (string_drop _s6329_ _s6330_) with
+ | _s6331_ =>
+ match (hex_bits_7_matches_prefix _s6331_) with
+ | Some (v__1382, existT _ _s6332_ _) =>
+ if eq_vec (subrange_vec_dec v__1382 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1382 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1382 6 2 in
+ match (string_drop _s6331_ _s6332_) with
+ | s_ => Some (rsc1, rsc2, uimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5 * string))).
+
+Definition _s6302_ (_s6303_ : string) : M (option ((mword 3 * mword 3 * mword 5 * string))) :=
+ let _s6304_ := _s6303_ in
+ (if string_startswith _s6304_ "c.flw" then
+ (match (string_drop _s6304_ (projT1 (string_length "c.flw"))) with
+ | _s6305_ =>
+ (spc_matches_prefix _s6305_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s6306_ _) =>
+ (match (string_drop _s6305_ _s6306_) with
+ | _s6307_ =>
+ (creg_name_matches_prefix _s6307_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rdc, existT _ _s6308_ _) =>
+ (match (string_drop _s6307_ _s6308_) with
+ | _s6309_ =>
+ (sep_matches_prefix _s6309_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (tt, existT _ _s6310_ _) =>
+ (match (string_drop _s6309_ _s6310_) with
+ | _s6311_ =>
+ (creg_name_matches_prefix _s6311_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (rsc, existT _ _s6312_ _) =>
+ (match (string_drop _s6311_ _s6312_) with
+ | _s6313_ =>
+ (sep_matches_prefix _s6313_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (tt, existT _ _s6314_ _) =>
+ match (string_drop _s6313_ _s6314_) with
+ | _s6315_ =>
+ match (hex_bits_7_matches_prefix _s6315_) with
+ | Some (v__1384, existT _ _s6316_ _) =>
+ if eq_vec (subrange_vec_dec v__1384 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1384 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1384 6 2 in
+ match (string_drop _s6315_ _s6316_) with
+ | s_ => Some (rdc, rsc, uimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ end)
+ : M (option ((mword 3 * mword 3 * mword 5 * string)))
+ else returnm None)
+ : M (option ((mword 3 * mword 3 * mword 5 * string))).
+
+Definition _s6290_ (_s6291_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s6292_ := _s6291_ in
+ (if string_startswith _s6292_ "c.fswsp" then
+ (match (string_drop _s6292_ (projT1 (string_length "c.fswsp"))) with
+ | _s6293_ =>
+ (spc_matches_prefix _s6293_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s6294_ _) =>
+ (match (string_drop _s6293_ _s6294_) with
+ | _s6295_ =>
+ (reg_name_matches_prefix _s6295_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s6296_ _) =>
+ (match (string_drop _s6295_ _s6296_) with
+ | _s6297_ =>
+ (sep_matches_prefix _s6297_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s6298_ _) =>
+ match (string_drop _s6297_ _s6298_) with
+ | _s6299_ =>
+ match (hex_bits_6_matches_prefix _s6299_) with
+ | Some (uimm, existT _ _s6300_ _) =>
+ match (string_drop _s6299_ _s6300_) with
+ | s_ => Some (rd, uimm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6 * string))).
+
+Definition _s6278_ (_s6279_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s6280_ := _s6279_ in
+ (if string_startswith _s6280_ "c.flwsp" then
+ (match (string_drop _s6280_ (projT1 (string_length "c.flwsp"))) with
+ | _s6281_ =>
+ (spc_matches_prefix _s6281_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ (match w__0 with
+ | Some (tt, existT _ _s6282_ _) =>
+ (match (string_drop _s6281_ _s6282_) with
+ | _s6283_ =>
+ (reg_name_matches_prefix _s6283_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (rd, existT _ _s6284_ _) =>
+ (match (string_drop _s6283_ _s6284_) with
+ | _s6285_ =>
+ (sep_matches_prefix _s6285_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__2 with
+ | Some (tt, existT _ _s6286_ _) =>
+ match (string_drop _s6285_ _s6286_) with
+ | _s6287_ =>
+ match (hex_bits_6_matches_prefix _s6287_) with
+ | Some (imm, existT _ _s6288_ _) =>
+ match (string_drop _s6287_ _s6288_) with
+ | s_ => Some (rd, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ end)
+ : M (option ((mword 5 * mword 6 * string)))
+ else returnm None)
+ : M (option ((mword 5 * mword 6 * string))).
+
+Definition _s6265_ (_s6266_ : string) : M (option ((f_un_op_S * mword 5 * mword 5 * string))) :=
+ (match _s6266_ with
+ | _s6267_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s6267_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCLASS_S, existT _ _s6268_ _) =>
+ (match (string_drop _s6267_ _s6268_) with
+ | _s6269_ =>
+ (spc_matches_prefix _s6269_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6270_ _) =>
+ (match (string_drop _s6269_ _s6270_) with
+ | _s6271_ =>
+ (reg_name_matches_prefix _s6271_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6272_ _) =>
+ (match (string_drop _s6271_ _s6272_) with
+ | _s6273_ =>
+ (sep_matches_prefix _s6273_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6274_ _) =>
+ (match (string_drop _s6273_ _s6274_) with
+ | _s6275_ =>
+ (freg_name_matches_prefix _s6275_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s6276_ _) =>
+ match (string_drop _s6275_ _s6276_) with
+ | s_ => Some (FCLASS_S, rd, rs1, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string))).
+
+Definition _s6252_ (_s6253_ : string) : M (option ((f_un_op_S * mword 5 * mword 5 * string))) :=
+ (match _s6253_ with
+ | _s6254_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s6254_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_W_X, existT _ _s6255_ _) =>
+ (match (string_drop _s6254_ _s6255_) with
+ | _s6256_ =>
+ (spc_matches_prefix _s6256_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6257_ _) =>
+ (match (string_drop _s6256_ _s6257_) with
+ | _s6258_ =>
+ (freg_name_matches_prefix _s6258_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6259_ _) =>
+ (match (string_drop _s6258_ _s6259_) with
+ | _s6260_ =>
+ (sep_matches_prefix _s6260_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6261_ _) =>
+ (match (string_drop _s6260_ _s6261_) with
+ | _s6262_ =>
+ (reg_name_matches_prefix _s6262_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s6263_ _) =>
+ match (string_drop _s6262_ _s6263_) with
+ | s_ => Some (FMV_W_X, rd, rs1, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string))).
+
+Definition _s6239_ (_s6240_ : string) : M (option ((f_un_op_S * mword 5 * mword 5 * string))) :=
+ (match _s6240_ with
+ | _s6241_ =>
+ (f_un_type_mnemonic_S_matches_prefix _s6241_) >>= fun w__0 : option ((f_un_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMV_X_W, existT _ _s6242_ _) =>
+ (match (string_drop _s6241_ _s6242_) with
+ | _s6243_ =>
+ (spc_matches_prefix _s6243_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6244_ _) =>
+ (match (string_drop _s6243_ _s6244_) with
+ | _s6245_ =>
+ (reg_name_matches_prefix _s6245_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6246_ _) =>
+ (match (string_drop _s6245_ _s6246_) with
+ | _s6247_ =>
+ (sep_matches_prefix _s6247_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6248_ _) =>
+ (match (string_drop _s6247_ _s6248_) with
+ | _s6249_ =>
+ (freg_name_matches_prefix _s6249_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__4 with
+ | Some (rs1, existT _ _s6250_ _) =>
+ match (string_drop _s6249_ _s6250_) with
+ | s_ => Some (FMV_X_W, rd, rs1, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_un_op_S * mword 5 * mword 5 * string))).
+
+Definition _s6222_ (_s6223_ : string)
+: M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6223_ with
+ | _s6224_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s6224_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLE_S, existT _ _s6225_ _) =>
+ (match (string_drop _s6224_ _s6225_) with
+ | _s6226_ =>
+ (spc_matches_prefix _s6226_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6227_ _) =>
+ (match (string_drop _s6226_ _s6227_) with
+ | _s6228_ =>
+ (reg_name_matches_prefix _s6228_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6229_ _) =>
+ (match (string_drop _s6228_ _s6229_) with
+ | _s6230_ =>
+ (sep_matches_prefix _s6230_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6231_ _) =>
+ (match (string_drop _s6230_ _s6231_) with
+ | _s6232_ =>
+ (freg_name_matches_prefix _s6232_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6233_ _) =>
+ (match (string_drop _s6232_ _s6233_) with
+ | _s6234_ =>
+ (sep_matches_prefix _s6234_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6235_ _) =>
+ (match (string_drop _s6234_ _s6235_) with
+ | _s6236_ =>
+ (freg_name_matches_prefix _s6236_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6237_ _) =>
+ match (string_drop _s6236_
+ _s6237_) with
+ | s_ =>
+ Some (FLE_S, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6205_ (_s6206_ : string)
+: M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6206_ with
+ | _s6207_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s6207_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FLT_S, existT _ _s6208_ _) =>
+ (match (string_drop _s6207_ _s6208_) with
+ | _s6209_ =>
+ (spc_matches_prefix _s6209_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6210_ _) =>
+ (match (string_drop _s6209_ _s6210_) with
+ | _s6211_ =>
+ (reg_name_matches_prefix _s6211_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6212_ _) =>
+ (match (string_drop _s6211_ _s6212_) with
+ | _s6213_ =>
+ (sep_matches_prefix _s6213_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6214_ _) =>
+ (match (string_drop _s6213_ _s6214_) with
+ | _s6215_ =>
+ (freg_name_matches_prefix _s6215_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6216_ _) =>
+ (match (string_drop _s6215_ _s6216_) with
+ | _s6217_ =>
+ (sep_matches_prefix _s6217_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6218_ _) =>
+ (match (string_drop _s6217_ _s6218_) with
+ | _s6219_ =>
+ (freg_name_matches_prefix _s6219_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6220_ _) =>
+ match (string_drop _s6219_
+ _s6220_) with
+ | s_ =>
+ Some (FLT_S, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6188_ (_s6189_ : string)
+: M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6189_ with
+ | _s6190_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s6190_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FEQ_S, existT _ _s6191_ _) =>
+ (match (string_drop _s6190_ _s6191_) with
+ | _s6192_ =>
+ (spc_matches_prefix _s6192_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6193_ _) =>
+ (match (string_drop _s6192_ _s6193_) with
+ | _s6194_ =>
+ (reg_name_matches_prefix _s6194_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6195_ _) =>
+ (match (string_drop _s6194_ _s6195_) with
+ | _s6196_ =>
+ (sep_matches_prefix _s6196_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6197_ _) =>
+ (match (string_drop _s6196_ _s6197_) with
+ | _s6198_ =>
+ (freg_name_matches_prefix _s6198_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6199_ _) =>
+ (match (string_drop _s6198_ _s6199_) with
+ | _s6200_ =>
+ (sep_matches_prefix _s6200_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6201_ _) =>
+ (match (string_drop _s6200_ _s6201_) with
+ | _s6202_ =>
+ (freg_name_matches_prefix _s6202_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6203_ _) =>
+ match (string_drop _s6202_
+ _s6203_) with
+ | s_ =>
+ Some (FEQ_S, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6171_ (_s6172_ : string)
+: M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6172_ with
+ | _s6173_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s6173_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMAX_S, existT _ _s6174_ _) =>
+ (match (string_drop _s6173_ _s6174_) with
+ | _s6175_ =>
+ (spc_matches_prefix _s6175_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6176_ _) =>
+ (match (string_drop _s6175_ _s6176_) with
+ | _s6177_ =>
+ (freg_name_matches_prefix _s6177_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6178_ _) =>
+ (match (string_drop _s6177_ _s6178_) with
+ | _s6179_ =>
+ (sep_matches_prefix _s6179_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6180_ _) =>
+ (match (string_drop _s6179_ _s6180_) with
+ | _s6181_ =>
+ (freg_name_matches_prefix _s6181_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6182_ _) =>
+ (match (string_drop _s6181_ _s6182_) with
+ | _s6183_ =>
+ (sep_matches_prefix _s6183_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6184_ _) =>
+ (match (string_drop _s6183_ _s6184_) with
+ | _s6185_ =>
+ (freg_name_matches_prefix _s6185_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6186_ _) =>
+ match (string_drop _s6185_
+ _s6186_) with
+ | s_ =>
+ Some
+ (FMAX_S, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6154_ (_s6155_ : string)
+: M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6155_ with
+ | _s6156_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s6156_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FMIN_S, existT _ _s6157_ _) =>
+ (match (string_drop _s6156_ _s6157_) with
+ | _s6158_ =>
+ (spc_matches_prefix _s6158_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6159_ _) =>
+ (match (string_drop _s6158_ _s6159_) with
+ | _s6160_ =>
+ (freg_name_matches_prefix _s6160_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6161_ _) =>
+ (match (string_drop _s6160_ _s6161_) with
+ | _s6162_ =>
+ (sep_matches_prefix _s6162_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6163_ _) =>
+ (match (string_drop _s6162_ _s6163_) with
+ | _s6164_ =>
+ (freg_name_matches_prefix _s6164_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6165_ _) =>
+ (match (string_drop _s6164_ _s6165_) with
+ | _s6166_ =>
+ (sep_matches_prefix _s6166_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6167_ _) =>
+ (match (string_drop _s6166_ _s6167_) with
+ | _s6168_ =>
+ (freg_name_matches_prefix _s6168_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6169_ _) =>
+ match (string_drop _s6168_
+ _s6169_) with
+ | s_ =>
+ Some
+ (FMIN_S, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6137_ (_s6138_ : string)
+: M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6138_ with
+ | _s6139_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s6139_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJX_S, existT _ _s6140_ _) =>
+ (match (string_drop _s6139_ _s6140_) with
+ | _s6141_ =>
+ (spc_matches_prefix _s6141_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6142_ _) =>
+ (match (string_drop _s6141_ _s6142_) with
+ | _s6143_ =>
+ (freg_name_matches_prefix _s6143_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6144_ _) =>
+ (match (string_drop _s6143_ _s6144_) with
+ | _s6145_ =>
+ (sep_matches_prefix _s6145_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6146_ _) =>
+ (match (string_drop _s6145_ _s6146_) with
+ | _s6147_ =>
+ (freg_name_matches_prefix _s6147_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6148_ _) =>
+ (match (string_drop _s6147_ _s6148_) with
+ | _s6149_ =>
+ (sep_matches_prefix _s6149_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6150_ _) =>
+ (match (string_drop _s6149_ _s6150_) with
+ | _s6151_ =>
+ (freg_name_matches_prefix _s6151_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6152_ _) =>
+ match (string_drop _s6151_
+ _s6152_) with
+ | s_ =>
+ Some
+ (FSGNJX_S, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6120_ (_s6121_ : string)
+: M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6121_ with
+ | _s6122_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s6122_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJN_S, existT _ _s6123_ _) =>
+ (match (string_drop _s6122_ _s6123_) with
+ | _s6124_ =>
+ (spc_matches_prefix _s6124_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6125_ _) =>
+ (match (string_drop _s6124_ _s6125_) with
+ | _s6126_ =>
+ (freg_name_matches_prefix _s6126_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6127_ _) =>
+ (match (string_drop _s6126_ _s6127_) with
+ | _s6128_ =>
+ (sep_matches_prefix _s6128_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6129_ _) =>
+ (match (string_drop _s6128_ _s6129_) with
+ | _s6130_ =>
+ (freg_name_matches_prefix _s6130_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6131_ _) =>
+ (match (string_drop _s6130_ _s6131_) with
+ | _s6132_ =>
+ (sep_matches_prefix _s6132_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6133_ _) =>
+ (match (string_drop _s6132_ _s6133_) with
+ | _s6134_ =>
+ (freg_name_matches_prefix _s6134_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6135_ _) =>
+ match (string_drop _s6134_
+ _s6135_) with
+ | s_ =>
+ Some
+ (FSGNJN_S, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6103_ (_s6104_ : string)
+: M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s6104_ with
+ | _s6105_ =>
+ (f_bin_type_mnemonic_S_matches_prefix _s6105_) >>= fun w__0 : option ((f_bin_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSGNJ_S, existT _ _s6106_ _) =>
+ (match (string_drop _s6105_ _s6106_) with
+ | _s6107_ =>
+ (spc_matches_prefix _s6107_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6108_ _) =>
+ (match (string_drop _s6107_ _s6108_) with
+ | _s6109_ =>
+ (freg_name_matches_prefix _s6109_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6110_ _) =>
+ (match (string_drop _s6109_ _s6110_) with
+ | _s6111_ =>
+ (sep_matches_prefix _s6111_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6112_ _) =>
+ (match (string_drop _s6111_ _s6112_) with
+ | _s6113_ =>
+ (freg_name_matches_prefix _s6113_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6114_ _) =>
+ (match (string_drop _s6113_ _s6114_) with
+ | _s6115_ =>
+ (sep_matches_prefix _s6115_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6116_ _) =>
+ (match (string_drop _s6115_ _s6116_) with
+ | _s6117_ =>
+ (freg_name_matches_prefix _s6117_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s6118_ _) =>
+ match (string_drop _s6117_
+ _s6118_) with
+ | s_ =>
+ Some
+ (FSGNJ_S, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)))
+ end)
+ : M (option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string))).
+
+Definition _s6086_ (_s6087_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6087_ with
+ | _s6088_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s6088_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_LU, existT _ _s6089_ _) =>
+ (match (string_drop _s6088_ _s6089_) with
+ | _s6090_ =>
+ (spc_matches_prefix _s6090_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6091_ _) =>
+ (match (string_drop _s6090_ _s6091_) with
+ | _s6092_ =>
+ (freg_name_matches_prefix _s6092_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6093_ _) =>
+ (match (string_drop _s6092_ _s6093_) with
+ | _s6094_ =>
+ (sep_matches_prefix _s6094_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6095_ _) =>
+ (match (string_drop _s6094_ _s6095_) with
+ | _s6096_ =>
+ (reg_name_matches_prefix _s6096_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6097_ _) =>
+ (match (string_drop _s6096_ _s6097_) with
+ | _s6098_ =>
+ (sep_matches_prefix _s6098_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6099_ _) =>
+ (match (string_drop _s6098_ _s6099_) with
+ | _s6100_ =>
+ (frm_mnemonic_matches_prefix _s6100_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6101_ _) =>
+ match (string_drop _s6100_
+ _s6101_) with
+ | s_ =>
+ Some
+ (FCVT_S_LU, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6069_ (_s6070_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6070_ with
+ | _s6071_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s6071_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_L, existT _ _s6072_ _) =>
+ (match (string_drop _s6071_ _s6072_) with
+ | _s6073_ =>
+ (spc_matches_prefix _s6073_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6074_ _) =>
+ (match (string_drop _s6073_ _s6074_) with
+ | _s6075_ =>
+ (freg_name_matches_prefix _s6075_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6076_ _) =>
+ (match (string_drop _s6075_ _s6076_) with
+ | _s6077_ =>
+ (sep_matches_prefix _s6077_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6078_ _) =>
+ (match (string_drop _s6077_ _s6078_) with
+ | _s6079_ =>
+ (reg_name_matches_prefix _s6079_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6080_ _) =>
+ (match (string_drop _s6079_ _s6080_) with
+ | _s6081_ =>
+ (sep_matches_prefix _s6081_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6082_ _) =>
+ (match (string_drop _s6081_ _s6082_) with
+ | _s6083_ =>
+ (frm_mnemonic_matches_prefix _s6083_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6084_ _) =>
+ match (string_drop _s6083_
+ _s6084_) with
+ | s_ =>
+ Some
+ (FCVT_S_L, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6052_ (_s6053_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6053_ with
+ | _s6054_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s6054_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_LU_S, existT _ _s6055_ _) =>
+ (match (string_drop _s6054_ _s6055_) with
+ | _s6056_ =>
+ (spc_matches_prefix _s6056_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6057_ _) =>
+ (match (string_drop _s6056_ _s6057_) with
+ | _s6058_ =>
+ (reg_name_matches_prefix _s6058_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6059_ _) =>
+ (match (string_drop _s6058_ _s6059_) with
+ | _s6060_ =>
+ (sep_matches_prefix _s6060_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6061_ _) =>
+ (match (string_drop _s6060_ _s6061_) with
+ | _s6062_ =>
+ (freg_name_matches_prefix _s6062_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6063_ _) =>
+ (match (string_drop _s6062_ _s6063_) with
+ | _s6064_ =>
+ (sep_matches_prefix _s6064_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6065_ _) =>
+ (match (string_drop _s6064_ _s6065_) with
+ | _s6066_ =>
+ (frm_mnemonic_matches_prefix _s6066_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6067_ _) =>
+ match (string_drop _s6066_
+ _s6067_) with
+ | s_ =>
+ Some
+ (FCVT_LU_S, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6035_ (_s6036_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6036_ with
+ | _s6037_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s6037_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_L_S, existT _ _s6038_ _) =>
+ (match (string_drop _s6037_ _s6038_) with
+ | _s6039_ =>
+ (spc_matches_prefix _s6039_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6040_ _) =>
+ (match (string_drop _s6039_ _s6040_) with
+ | _s6041_ =>
+ (reg_name_matches_prefix _s6041_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6042_ _) =>
+ (match (string_drop _s6041_ _s6042_) with
+ | _s6043_ =>
+ (sep_matches_prefix _s6043_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6044_ _) =>
+ (match (string_drop _s6043_ _s6044_) with
+ | _s6045_ =>
+ (freg_name_matches_prefix _s6045_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6046_ _) =>
+ (match (string_drop _s6045_ _s6046_) with
+ | _s6047_ =>
+ (sep_matches_prefix _s6047_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6048_ _) =>
+ (match (string_drop _s6047_ _s6048_) with
+ | _s6049_ =>
+ (frm_mnemonic_matches_prefix _s6049_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6050_ _) =>
+ match (string_drop _s6049_
+ _s6050_) with
+ | s_ =>
+ Some
+ (FCVT_L_S, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6018_ (_s6019_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6019_ with
+ | _s6020_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s6020_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_WU, existT _ _s6021_ _) =>
+ (match (string_drop _s6020_ _s6021_) with
+ | _s6022_ =>
+ (spc_matches_prefix _s6022_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6023_ _) =>
+ (match (string_drop _s6022_ _s6023_) with
+ | _s6024_ =>
+ (freg_name_matches_prefix _s6024_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6025_ _) =>
+ (match (string_drop _s6024_ _s6025_) with
+ | _s6026_ =>
+ (sep_matches_prefix _s6026_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6027_ _) =>
+ (match (string_drop _s6026_ _s6027_) with
+ | _s6028_ =>
+ (reg_name_matches_prefix _s6028_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6029_ _) =>
+ (match (string_drop _s6028_ _s6029_) with
+ | _s6030_ =>
+ (sep_matches_prefix _s6030_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6031_ _) =>
+ (match (string_drop _s6030_ _s6031_) with
+ | _s6032_ =>
+ (frm_mnemonic_matches_prefix _s6032_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6033_ _) =>
+ match (string_drop _s6032_
+ _s6033_) with
+ | s_ =>
+ Some
+ (FCVT_S_WU, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s6001_ (_s6002_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s6002_ with
+ | _s6003_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s6003_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_S_W, existT _ _s6004_ _) =>
+ (match (string_drop _s6003_ _s6004_) with
+ | _s6005_ =>
+ (spc_matches_prefix _s6005_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s6006_ _) =>
+ (match (string_drop _s6005_ _s6006_) with
+ | _s6007_ =>
+ (freg_name_matches_prefix _s6007_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s6008_ _) =>
+ (match (string_drop _s6007_ _s6008_) with
+ | _s6009_ =>
+ (sep_matches_prefix _s6009_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s6010_ _) =>
+ (match (string_drop _s6009_ _s6010_) with
+ | _s6011_ =>
+ (reg_name_matches_prefix _s6011_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s6012_ _) =>
+ (match (string_drop _s6011_ _s6012_) with
+ | _s6013_ =>
+ (sep_matches_prefix _s6013_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s6014_ _) =>
+ (match (string_drop _s6013_ _s6014_) with
+ | _s6015_ =>
+ (frm_mnemonic_matches_prefix _s6015_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s6016_ _) =>
+ match (string_drop _s6015_
+ _s6016_) with
+ | s_ =>
+ Some
+ (FCVT_S_W, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s5984_ (_s5985_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s5985_ with
+ | _s5986_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s5986_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_WU_S, existT _ _s5987_ _) =>
+ (match (string_drop _s5986_ _s5987_) with
+ | _s5988_ =>
+ (spc_matches_prefix _s5988_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s5989_ _) =>
+ (match (string_drop _s5988_ _s5989_) with
+ | _s5990_ =>
+ (reg_name_matches_prefix _s5990_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s5991_ _) =>
+ (match (string_drop _s5990_ _s5991_) with
+ | _s5992_ =>
+ (sep_matches_prefix _s5992_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s5993_ _) =>
+ (match (string_drop _s5992_ _s5993_) with
+ | _s5994_ =>
+ (freg_name_matches_prefix _s5994_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s5995_ _) =>
+ (match (string_drop _s5994_ _s5995_) with
+ | _s5996_ =>
+ (sep_matches_prefix _s5996_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s5997_ _) =>
+ (match (string_drop _s5996_ _s5997_) with
+ | _s5998_ =>
+ (frm_mnemonic_matches_prefix _s5998_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s5999_ _) =>
+ match (string_drop _s5998_
+ _s5999_) with
+ | s_ =>
+ Some
+ (FCVT_WU_S, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s5967_ (_s5968_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s5968_ with
+ | _s5969_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s5969_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FCVT_W_S, existT _ _s5970_ _) =>
+ (match (string_drop _s5969_ _s5970_) with
+ | _s5971_ =>
+ (spc_matches_prefix _s5971_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s5972_ _) =>
+ (match (string_drop _s5971_ _s5972_) with
+ | _s5973_ =>
+ (reg_name_matches_prefix _s5973_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s5974_ _) =>
+ (match (string_drop _s5973_ _s5974_) with
+ | _s5975_ =>
+ (sep_matches_prefix _s5975_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s5976_ _) =>
+ (match (string_drop _s5975_ _s5976_) with
+ | _s5977_ =>
+ (freg_name_matches_prefix _s5977_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s5978_ _) =>
+ (match (string_drop _s5977_ _s5978_) with
+ | _s5979_ =>
+ (sep_matches_prefix _s5979_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s5980_ _) =>
+ (match (string_drop _s5979_ _s5980_) with
+ | _s5981_ =>
+ (frm_mnemonic_matches_prefix _s5981_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s5982_ _) =>
+ match (string_drop _s5981_
+ _s5982_) with
+ | s_ =>
+ Some
+ (FCVT_W_S, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s5950_ (_s5951_ : string)
+: M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s5951_ with
+ | _s5952_ =>
+ (f_un_rm_type_mnemonic_S_matches_prefix _s5952_) >>= fun w__0 : option ((f_un_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (FSQRT_S, existT _ _s5953_ _) =>
+ (match (string_drop _s5952_ _s5953_) with
+ | _s5954_ =>
+ (spc_matches_prefix _s5954_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s5955_ _) =>
+ (match (string_drop _s5954_ _s5955_) with
+ | _s5956_ =>
+ (freg_name_matches_prefix _s5956_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s5957_ _) =>
+ (match (string_drop _s5956_ _s5957_) with
+ | _s5958_ =>
+ (sep_matches_prefix _s5958_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s5959_ _) =>
+ (match (string_drop _s5958_ _s5959_) with
+ | _s5960_ =>
+ (freg_name_matches_prefix _s5960_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s5961_ _) =>
+ (match (string_drop _s5960_ _s5961_) with
+ | _s5962_ =>
+ (sep_matches_prefix _s5962_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s5963_ _) =>
+ (match (string_drop _s5962_ _s5963_) with
+ | _s5964_ =>
+ (frm_mnemonic_matches_prefix _s5964_) >>= fun w__6 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rm, existT _ _s5965_ _) =>
+ match (string_drop _s5964_
+ _s5965_) with
+ | s_ =>
+ Some
+ (FSQRT_S, rd, rs1, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s5929_ (_s5930_ : string)
+: M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s5930_ with
+ | _s5931_ =>
+ (f_bin_rm_type_mnemonic_S_matches_prefix _s5931_) >>= fun w__0 : option ((f_bin_rm_op_S * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s5932_ _) =>
+ (match (string_drop _s5931_ _s5932_) with
+ | _s5933_ =>
+ (spc_matches_prefix _s5933_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s5934_ _) =>
+ (match (string_drop _s5933_ _s5934_) with
+ | _s5935_ =>
+ (freg_name_matches_prefix _s5935_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s5936_ _) =>
+ (match (string_drop _s5935_ _s5936_) with
+ | _s5937_ =>
+ (sep_matches_prefix _s5937_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s5938_ _) =>
+ (match (string_drop _s5937_ _s5938_) with
+ | _s5939_ =>
+ (freg_name_matches_prefix _s5939_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s5940_ _) =>
+ (match (string_drop _s5939_ _s5940_) with
+ | _s5941_ =>
+ (sep_matches_prefix _s5941_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s5942_ _) =>
+ (match (string_drop _s5941_ _s5942_) with
+ | _s5943_ =>
+ (freg_name_matches_prefix _s5943_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s5944_ _) =>
+ (match (string_drop _s5943_ _s5944_) with
+ | _s5945_ =>
+ (sep_matches_prefix _s5945_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s5946_ _) =>
+ (match (string_drop _s5945_
+ _s5946_) with
+ | _s5947_ =>
+ (frm_mnemonic_matches_prefix
+ _s5947_) >>= fun w__8 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__8 with
+ | Some
+ (rm, existT _ _s5948_ _) =>
+ match (string_drop
+ _s5947_
+ _s5948_) with
+ | s_ =>
+ Some
+ (op, rd, rs1, rs2, rm, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string))).
+ : M (option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string))).
-Definition _s3410_ (_s3411_ : string)
-: M (option ((csrop * mword 5 * mword 5 * mword 12 * string))) :=
-
- (match _s3411_ with
- | _s3412_ =>
- (csr_mnemonic_matches_prefix _s3412_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=
+Definition _s5904_ (_s5905_ : string)
+: M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string))) :=
+ (match _s5905_ with
+ | _s5906_ =>
+ (f_madd_type_mnemonic_S_matches_prefix _s5906_) >>= fun w__0 : option ((f_madd_op_S * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s3413_ _)) =>
- let _s3414_ := string_drop _s3412_ _s3413_ in
- (if ((string_startswith _s3414_ "i")) then
- (match (string_drop _s3414_ (projT1 (string_length "i"))) with
- | _s3415_ =>
- (spc_matches_prefix _s3415_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5907_ _) =>
+ (match (string_drop _s5906_ _s5907_) with
+ | _s5908_ =>
+ (spc_matches_prefix _s5908_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s5909_ _) =>
+ (match (string_drop _s5908_ _s5909_) with
+ | _s5910_ =>
+ (freg_name_matches_prefix _s5910_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s5911_ _) =>
+ (match (string_drop _s5910_ _s5911_) with
+ | _s5912_ =>
+ (sep_matches_prefix _s5912_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s5913_ _) =>
+ (match (string_drop _s5912_ _s5913_) with
+ | _s5914_ =>
+ (freg_name_matches_prefix _s5914_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (rs1, existT _ _s5915_ _) =>
+ (match (string_drop _s5914_ _s5915_) with
+ | _s5916_ =>
+ (sep_matches_prefix _s5916_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s5917_ _) =>
+ (match (string_drop _s5916_ _s5917_) with
+ | _s5918_ =>
+ (freg_name_matches_prefix _s5918_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some (rs2, existT _ _s5919_ _) =>
+ (match (string_drop _s5918_ _s5919_) with
+ | _s5920_ =>
+ (sep_matches_prefix _s5920_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__7 with
+ | Some (tt, existT _ _s5921_ _) =>
+ (match (string_drop _s5920_
+ _s5921_) with
+ | _s5922_ =>
+ (freg_name_matches_prefix
+ _s5922_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__8 with
+ | Some
+ (rs3, existT _ _s5923_ _) =>
+ (match (string_drop
+ _s5922_ _s5923_) with
+ | _s5924_ =>
+ (sep_matches_prefix
+ _s5924_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__9 with
+ | Some
+ (tt, existT _ _s5925_ _) =>
+ (match (string_drop
+ _s5924_
+ _s5925_) with
+ | _s5926_ =>
+ (frm_mnemonic_matches_prefix
+ _s5926_) >>= fun w__10 : option ((rounding_mode * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__10 with
+ | Some
+ (rm, existT _ _s5927_ _) =>
+ match (string_drop
+ _s5926_
+ _s5927_) with
+ | s_ =>
+ Some
+ (op, rd, rs1, rs2, rs3, rm, s_)
+ end
+ | _ =>
+ None
+ end)
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ | _ => returnm None
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)))
+ end)
+ : M (option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string))).
+
+Definition _s5880_ (_s5881_ : string)
+: M (option ((word_width * mword 5 * mword 12 * mword 5 * string))) :=
+ let _s5882_ := _s5881_ in
+ (if string_startswith _s5882_ "fs" then
+ (match (string_drop _s5882_ (projT1 (string_length "fs"))) with
+ | _s5883_ =>
+ (size_mnemonic_matches_prefix _s5883_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (width, existT _ _s5884_ _) =>
+ (match (string_drop _s5883_ _s5884_) with
+ | _s5885_ =>
+ (spc_matches_prefix _s5885_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s3416_ _)) =>
- (match (string_drop _s3415_ _s3416_) with
- | _s3417_ =>
- (reg_name_matches_prefix _s3417_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5886_ _) =>
+ (match (string_drop _s5885_ _s5886_) with
+ | _s5887_ =>
+ (freg_name_matches_prefix _s5887_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s3418_ _)) =>
- (match (string_drop _s3417_ _s3418_) with
- | _s3419_ =>
- (sep_matches_prefix _s3419_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s5888_ _) =>
+ (match (string_drop _s5887_ _s5888_) with
+ | _s5889_ =>
+ (sep_matches_prefix _s5889_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s3420_ _)) =>
- (match (string_drop _s3419_ _s3420_) with
- | _s3421_ =>
- (match (hex_bits_5_matches_prefix _s3421_) with
- | Some ((rs1, existT _ _s3422_ _)) =>
- (match (string_drop _s3421_ _s3422_) with
- | _s3423_ =>
- (sep_matches_prefix _s3423_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5890_ _) =>
+ (match (string_drop _s5889_ _s5890_) with
+ | _s5891_ =>
+ (match (hex_bits_12_matches_prefix _s5891_) with
+ | Some (imm, existT _ _s5892_ _) =>
+ (match (string_drop _s5891_ _s5892_) with
+ | _s5893_ =>
+ (opt_spc_matches_prefix _s5893_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s3424_ _)) =>
- (match (string_drop _s3423_ _s3424_) with
- | _s3425_ =>
- (csr_name_map_matches_prefix _s3425_) >>= fun w__5 : option ((mword 12 * {n : Z & ArithFact (n >=
- 0)})) =>
- returnm ((match w__5 with
- | Some
- ((csr, existT _ _s3426_ _)) =>
- match (string_drop _s3425_
- _s3426_) with
- | s_ =>
- Some
- ((op, rd, rs1, csr, s_))
- end
- | _ => None
- end)
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ | Some (tt, existT _ _s5894_ _) =>
+ let _s5895_ := string_drop _s5893_ _s5894_ in
+ (if string_startswith _s5895_ "(" then
+ (match (string_drop _s5895_
+ (projT1
+ (string_length "("))) with
+ | _s5896_ =>
+ (opt_spc_matches_prefix _s5896_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s5897_ _) =>
+ (match (string_drop _s5896_ _s5897_) with
+ | _s5898_ =>
+ (reg_name_matches_prefix
+ _s5898_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some
+ (rs1, existT _ _s5899_ _) =>
+ (match (string_drop _s5898_
+ _s5899_) with
+ | _s5900_ =>
+ (opt_spc_matches_prefix
+ _s5900_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__7 with
+ | Some
+ (tt, existT _ _s5901_ _) =>
+ let _s5902_ :=
+ string_drop
+ _s5900_
+ _s5901_ in
+ if string_startswith
+ _s5902_
+ ")"
+ then
+ match (string_drop
+ _s5902_
+ (projT1
+ (string_length
+ ")"))) with
+ | s_ =>
+ Some
+ (width, rs2, imm, rs1, s_)
+ end
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string))).
+
+Definition _s5856_ (_s5857_ : string)
+: M (option ((word_width * mword 5 * mword 12 * mword 5 * string))) :=
+ let _s5858_ := _s5857_ in
+ (if string_startswith _s5858_ "fl" then
+ (match (string_drop _s5858_ (projT1 (string_length "fl"))) with
+ | _s5859_ =>
+ (size_mnemonic_matches_prefix _s5859_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (width, existT _ _s5860_ _) =>
+ (match (string_drop _s5859_ _s5860_) with
+ | _s5861_ =>
+ (spc_matches_prefix _s5861_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s5862_ _) =>
+ (match (string_drop _s5861_ _s5862_) with
+ | _s5863_ =>
+ (freg_name_matches_prefix _s5863_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s5864_ _) =>
+ (match (string_drop _s5863_ _s5864_) with
+ | _s5865_ =>
+ (sep_matches_prefix _s5865_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s5866_ _) =>
+ (match (string_drop _s5865_ _s5866_) with
+ | _s5867_ =>
+ (match (hex_bits_12_matches_prefix _s5867_) with
+ | Some (imm, existT _ _s5868_ _) =>
+ (match (string_drop _s5867_ _s5868_) with
+ | _s5869_ =>
+ (opt_spc_matches_prefix _s5869_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (tt, existT _ _s5870_ _) =>
+ let _s5871_ := string_drop _s5869_ _s5870_ in
+ (if string_startswith _s5871_ "(" then
+ (match (string_drop _s5871_
+ (projT1
+ (string_length "("))) with
+ | _s5872_ =>
+ (opt_spc_matches_prefix _s5872_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s5873_ _) =>
+ (match (string_drop _s5872_ _s5873_) with
+ | _s5874_ =>
+ (reg_name_matches_prefix
+ _s5874_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__6 with
+ | Some
+ (rs1, existT _ _s5875_ _) =>
+ (match (string_drop _s5874_
+ _s5875_) with
+ | _s5876_ =>
+ (opt_spc_matches_prefix
+ _s5876_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__7 with
+ | Some
+ (tt, existT _ _s5877_ _) =>
+ let _s5878_ :=
+ string_drop
+ _s5876_
+ _s5877_ in
+ if string_startswith
+ _s5878_
+ ")"
+ then
+ match (string_drop
+ _s5878_
+ (projT1
+ (string_length
+ ")"))) with
+ | s_ =>
+ Some
+ (width, rd, imm, rs1, s_)
+ end
+ else None
+ | _ => None
+ end)
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- else returnm (None : option ((csrop * mword 5 * mword 5 * mword 12 * string))))
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((csrop * mword 5 * mword 5 * mword 12 * string)))
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string)))
+ else returnm None)
+ : M (option ((word_width * mword 5 * mword 12 * mword 5 * string))).
+
+Definition _s5852_ (_s5853_ : string) : option string :=
+ let _s5854_ := _s5853_ in
+ if string_startswith _s5854_ "uret" then
+ match (string_drop _s5854_ (projT1 (string_length "uret"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s5835_ (_s5836_ : string) : M (option ((csrop * mword 5 * mword 12 * mword 5 * string))) :=
+ (match _s5836_ with
+ | _s5837_ =>
+ (csr_mnemonic_matches_prefix _s5837_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s5838_ _) =>
+ (match (string_drop _s5837_ _s5838_) with
+ | _s5839_ =>
+ (spc_matches_prefix _s5839_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s5840_ _) =>
+ (match (string_drop _s5839_ _s5840_) with
+ | _s5841_ =>
+ (reg_name_matches_prefix _s5841_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s5842_ _) =>
+ (match (string_drop _s5841_ _s5842_) with
+ | _s5843_ =>
+ (sep_matches_prefix _s5843_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s5844_ _) =>
+ (match (string_drop _s5843_ _s5844_) with
+ | _s5845_ =>
+ (csr_name_map_matches_prefix _s5845_) >>= fun w__4 : option ((mword 12 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (csr, existT _ _s5846_ _) =>
+ (match (string_drop _s5845_ _s5846_) with
+ | _s5847_ =>
+ (sep_matches_prefix _s5847_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__5 with
+ | Some (tt, existT _ _s5848_ _) =>
+ (match (string_drop _s5847_ _s5848_) with
+ | _s5849_ =>
+ (reg_name_matches_prefix _s5849_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__6 with
+ | Some (rs1, existT _ _s5850_ _) =>
+ match (string_drop _s5849_
+ _s5850_) with
+ | s_ =>
+ Some (op, rd, csr, rs1, s_)
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string)))
- end)
- : M (option ((csrop * mword 5 * mword 5 * mword 12 * string))).
-
-Definition _s3391_ (_s3392_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5 * string))) :=
-
- let _s3393_ := _s3392_ in
- (if ((string_startswith _s3393_ "rem")) then
- (match (string_drop _s3393_ (projT1 (string_length "rem"))) with
- | _s3394_ =>
- (maybe_not_u_matches_prefix _s3394_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string))).
+
+Definition _s5817_ (_s5818_ : string) : M (option ((csrop * mword 5 * mword 12 * mword 5 * string))) :=
+ (match _s5818_ with
+ | _s5819_ =>
+ (csr_mnemonic_matches_prefix _s5819_) >>= fun w__0 : option ((csrop * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__0 with
+ | Some (op, existT _ _s5820_ _) =>
+ let _s5821_ := string_drop _s5819_ _s5820_ in
+ (if string_startswith _s5821_ "i" then
+ (match (string_drop _s5821_ (projT1 (string_length "i"))) with
+ | _s5822_ =>
+ (spc_matches_prefix _s5822_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__1 with
+ | Some (tt, existT _ _s5823_ _) =>
+ (match (string_drop _s5822_ _s5823_) with
+ | _s5824_ =>
+ (reg_name_matches_prefix _s5824_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__2 with
+ | Some (rd, existT _ _s5825_ _) =>
+ (match (string_drop _s5824_ _s5825_) with
+ | _s5826_ =>
+ (sep_matches_prefix _s5826_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__3 with
+ | Some (tt, existT _ _s5827_ _) =>
+ (match (string_drop _s5826_ _s5827_) with
+ | _s5828_ =>
+ (csr_name_map_matches_prefix _s5828_) >>= fun w__4 : option ((mword 12 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ (match w__4 with
+ | Some (csr, existT _ _s5829_ _) =>
+ (match (string_drop _s5828_ _s5829_) with
+ | _s5830_ =>
+ (sep_matches_prefix _s5830_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__5 with
+ | Some (tt, existT _ _s5831_ _) =>
+ match (string_drop _s5830_ _s5831_) with
+ | _s5832_ =>
+ match (hex_bits_5_matches_prefix
+ _s5832_) with
+ | Some (rs1, existT _ _s5833_ _) =>
+ match (string_drop _s5832_
+ _s5833_) with
+ | s_ =>
+ Some (op, rd, csr, rs1, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ else returnm None)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string)))
+ end)
+ : M (option ((csrop * mword 5 * mword 12 * mword 5 * string))).
+
+Definition _s5798_ (_s5799_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5 * string))) :=
+ let _s5800_ := _s5799_ in
+ (if string_startswith _s5800_ "rem" then
+ (match (string_drop _s5800_ (projT1 (string_length "rem"))) with
+ | _s5801_ =>
+ (maybe_not_u_matches_prefix _s5801_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s3395_ _)) =>
- let _s3396_ := string_drop _s3394_ _s3395_ in
- (if ((string_startswith _s3396_ "w")) then
- (match (string_drop _s3396_ (projT1 (string_length "w"))) with
- | _s3397_ =>
- (spc_matches_prefix _s3397_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s5802_ _) =>
+ let _s5803_ := string_drop _s5801_ _s5802_ in
+ (if string_startswith _s5803_ "w" then
+ (match (string_drop _s5803_ (projT1 (string_length "w"))) with
+ | _s5804_ =>
+ (spc_matches_prefix _s5804_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s3398_ _)) =>
- (match (string_drop _s3397_ _s3398_) with
- | _s3399_ =>
- (reg_name_matches_prefix _s3399_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5805_ _) =>
+ (match (string_drop _s5804_ _s5805_) with
+ | _s5806_ =>
+ (reg_name_matches_prefix _s5806_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s3400_ _)) =>
- (match (string_drop _s3399_ _s3400_) with
- | _s3401_ =>
- (sep_matches_prefix _s3401_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5807_ _) =>
+ (match (string_drop _s5806_ _s5807_) with
+ | _s5808_ =>
+ (sep_matches_prefix _s5808_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s3402_ _)) =>
- (match (string_drop _s3401_ _s3402_) with
- | _s3403_ =>
- (reg_name_matches_prefix _s3403_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5809_ _) =>
+ (match (string_drop _s5808_ _s5809_) with
+ | _s5810_ =>
+ (reg_name_matches_prefix _s5810_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s3404_ _)) =>
- (match (string_drop _s3403_ _s3404_) with
- | _s3405_ =>
- (sep_matches_prefix _s3405_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5811_ _) =>
+ (match (string_drop _s5810_ _s5811_) with
+ | _s5812_ =>
+ (sep_matches_prefix _s5812_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s3406_ _)) =>
- (match (string_drop _s3405_ _s3406_) with
- | _s3407_ =>
- (reg_name_matches_prefix _s3407_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5813_ _) =>
+ (match (string_drop _s5812_ _s5813_) with
+ | _s5814_ =>
+ (reg_name_matches_prefix _s5814_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s3408_ _)) =>
- match (string_drop _s3407_
- _s3408_) with
- | s_ =>
- Some
- ((s, rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__6 with
+ | Some
+ (rs2, existT _ _s5815_ _) =>
+ match (string_drop _s5814_
+ _s5815_) with
+ | s_ =>
+ Some
+ (s, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string))).
-Definition _s3372_ (_s3373_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5 * string))) :=
-
- let _s3374_ := _s3373_ in
- (if ((string_startswith _s3374_ "div")) then
- (match (string_drop _s3374_ (projT1 (string_length "div"))) with
- | _s3375_ =>
- (maybe_not_u_matches_prefix _s3375_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s5779_ (_s5780_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5 * string))) :=
+ let _s5781_ := _s5780_ in
+ (if string_startswith _s5781_ "div" then
+ (match (string_drop _s5781_ (projT1 (string_length "div"))) with
+ | _s5782_ =>
+ (maybe_not_u_matches_prefix _s5782_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s3376_ _)) =>
- let _s3377_ := string_drop _s3375_ _s3376_ in
- (if ((string_startswith _s3377_ "w")) then
- (match (string_drop _s3377_ (projT1 (string_length "w"))) with
- | _s3378_ =>
- (spc_matches_prefix _s3378_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s5783_ _) =>
+ let _s5784_ := string_drop _s5782_ _s5783_ in
+ (if string_startswith _s5784_ "w" then
+ (match (string_drop _s5784_ (projT1 (string_length "w"))) with
+ | _s5785_ =>
+ (spc_matches_prefix _s5785_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s3379_ _)) =>
- (match (string_drop _s3378_ _s3379_) with
- | _s3380_ =>
- (reg_name_matches_prefix _s3380_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5786_ _) =>
+ (match (string_drop _s5785_ _s5786_) with
+ | _s5787_ =>
+ (reg_name_matches_prefix _s5787_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s3381_ _)) =>
- (match (string_drop _s3380_ _s3381_) with
- | _s3382_ =>
- (sep_matches_prefix _s3382_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5788_ _) =>
+ (match (string_drop _s5787_ _s5788_) with
+ | _s5789_ =>
+ (sep_matches_prefix _s5789_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s3383_ _)) =>
- (match (string_drop _s3382_ _s3383_) with
- | _s3384_ =>
- (reg_name_matches_prefix _s3384_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5790_ _) =>
+ (match (string_drop _s5789_ _s5790_) with
+ | _s5791_ =>
+ (reg_name_matches_prefix _s5791_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s3385_ _)) =>
- (match (string_drop _s3384_ _s3385_) with
- | _s3386_ =>
- (sep_matches_prefix _s3386_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5792_ _) =>
+ (match (string_drop _s5791_ _s5792_) with
+ | _s5793_ =>
+ (sep_matches_prefix _s5793_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s3387_ _)) =>
- (match (string_drop _s3386_ _s3387_) with
- | _s3388_ =>
- (reg_name_matches_prefix _s3388_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5794_ _) =>
+ (match (string_drop _s5793_ _s5794_) with
+ | _s5795_ =>
+ (reg_name_matches_prefix _s5795_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s3389_ _)) =>
- match (string_drop _s3388_
- _s3389_) with
- | s_ =>
- Some
- ((s, rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__6 with
+ | Some
+ (rs2, existT _ _s5796_ _) =>
+ match (string_drop _s5795_
+ _s5796_) with
+ | s_ =>
+ Some
+ (s, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string))).
-Definition _s3356_ (_s3357_ : string)
-: M (option ((mword 5 * mword 5 * mword 5 * string))) :=
-
- let _s3358_ := _s3357_ in
- (if ((string_startswith _s3358_ "mulw")) then
- (match (string_drop _s3358_ (projT1 (string_length "mulw"))) with
- | _s3359_ =>
- (spc_matches_prefix _s3359_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5763_ (_s5764_ : string) : M (option ((mword 5 * mword 5 * mword 5 * string))) :=
+ let _s5765_ := _s5764_ in
+ (if string_startswith _s5765_ "mulw" then
+ (match (string_drop _s5765_ (projT1 (string_length "mulw"))) with
+ | _s5766_ =>
+ (spc_matches_prefix _s5766_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3360_ _)) =>
- (match (string_drop _s3359_ _s3360_) with
- | _s3361_ =>
- (reg_name_matches_prefix _s3361_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5767_ _) =>
+ (match (string_drop _s5766_ _s5767_) with
+ | _s5768_ =>
+ (reg_name_matches_prefix _s5768_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s3362_ _)) =>
- (match (string_drop _s3361_ _s3362_) with
- | _s3363_ =>
- (sep_matches_prefix _s3363_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5769_ _) =>
+ (match (string_drop _s5768_ _s5769_) with
+ | _s5770_ =>
+ (sep_matches_prefix _s5770_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3364_ _)) =>
- (match (string_drop _s3363_ _s3364_) with
- | _s3365_ =>
- (reg_name_matches_prefix _s3365_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5771_ _) =>
+ (match (string_drop _s5770_ _s5771_) with
+ | _s5772_ =>
+ (reg_name_matches_prefix _s5772_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s3366_ _)) =>
- (match (string_drop _s3365_ _s3366_) with
- | _s3367_ =>
- (sep_matches_prefix _s3367_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5773_ _) =>
+ (match (string_drop _s5772_ _s5773_) with
+ | _s5774_ =>
+ (sep_matches_prefix _s5774_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s3368_ _)) =>
- (match (string_drop _s3367_ _s3368_) with
- | _s3369_ =>
- (reg_name_matches_prefix _s3369_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5775_ _) =>
+ (match (string_drop _s5774_ _s5775_) with
+ | _s5776_ =>
+ (reg_name_matches_prefix _s5776_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((rs2, existT _ _s3370_ _)) =>
- match (string_drop _s3369_ _s3370_) with
- | s_ => Some ((rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__5 with
+ | Some (rs2, existT _ _s5777_ _) =>
+ match (string_drop _s5776_ _s5777_) with
+ | s_ => Some (rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 5 * string)))
- else returnm (None : option ((mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 5 * string))).
-Definition _s3338_ (_s3339_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5 * string))) :=
-
- let _s3340_ := _s3339_ in
- (if ((string_startswith _s3340_ "rem")) then
- (match (string_drop _s3340_ (projT1 (string_length "rem"))) with
- | _s3341_ =>
- (maybe_not_u_matches_prefix _s3341_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s5745_ (_s5746_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5 * string))) :=
+ let _s5747_ := _s5746_ in
+ (if string_startswith _s5747_ "rem" then
+ (match (string_drop _s5747_ (projT1 (string_length "rem"))) with
+ | _s5748_ =>
+ (maybe_not_u_matches_prefix _s5748_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s3342_ _)) =>
- (match (string_drop _s3341_ _s3342_) with
- | _s3343_ =>
- (spc_matches_prefix _s3343_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s5749_ _) =>
+ (match (string_drop _s5748_ _s5749_) with
+ | _s5750_ =>
+ (spc_matches_prefix _s5750_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s3344_ _)) =>
- (match (string_drop _s3343_ _s3344_) with
- | _s3345_ =>
- (reg_name_matches_prefix _s3345_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5751_ _) =>
+ (match (string_drop _s5750_ _s5751_) with
+ | _s5752_ =>
+ (reg_name_matches_prefix _s5752_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s3346_ _)) =>
- (match (string_drop _s3345_ _s3346_) with
- | _s3347_ =>
- (sep_matches_prefix _s3347_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5753_ _) =>
+ (match (string_drop _s5752_ _s5753_) with
+ | _s5754_ =>
+ (sep_matches_prefix _s5754_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s3348_ _)) =>
- (match (string_drop _s3347_ _s3348_) with
- | _s3349_ =>
- (reg_name_matches_prefix _s3349_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5755_ _) =>
+ (match (string_drop _s5754_ _s5755_) with
+ | _s5756_ =>
+ (reg_name_matches_prefix _s5756_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s3350_ _)) =>
- (match (string_drop _s3349_ _s3350_) with
- | _s3351_ =>
- (sep_matches_prefix _s3351_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5757_ _) =>
+ (match (string_drop _s5756_ _s5757_) with
+ | _s5758_ =>
+ (sep_matches_prefix _s5758_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s3352_ _)) =>
- (match (string_drop _s3351_ _s3352_) with
- | _s3353_ =>
- (reg_name_matches_prefix _s3353_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5759_ _) =>
+ (match (string_drop _s5758_ _s5759_) with
+ | _s5760_ =>
+ (reg_name_matches_prefix _s5760_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s3354_ _)) =>
- match (string_drop _s3353_
- _s3354_) with
- | s_ =>
- Some
- ((s, rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s5761_ _) =>
+ match (string_drop _s5760_
+ _s5761_) with
+ | s_ =>
+ Some (s, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string))).
-Definition _s3320_ (_s3321_ : string)
-: M (option ((bool * mword 5 * mword 5 * mword 5 * string))) :=
-
- let _s3322_ := _s3321_ in
- (if ((string_startswith _s3322_ "div")) then
- (match (string_drop _s3322_ (projT1 (string_length "div"))) with
- | _s3323_ =>
- (maybe_not_u_matches_prefix _s3323_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=
+Definition _s5727_ (_s5728_ : string) : M (option ((bool * mword 5 * mword 5 * mword 5 * string))) :=
+ let _s5729_ := _s5728_ in
+ (if string_startswith _s5729_ "div" then
+ (match (string_drop _s5729_ (projT1 (string_length "div"))) with
+ | _s5730_ =>
+ (maybe_not_u_matches_prefix _s5730_) >>= fun w__0 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((s, existT _ _s3324_ _)) =>
- (match (string_drop _s3323_ _s3324_) with
- | _s3325_ =>
- (spc_matches_prefix _s3325_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (s, existT _ _s5731_ _) =>
+ (match (string_drop _s5730_ _s5731_) with
+ | _s5732_ =>
+ (spc_matches_prefix _s5732_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s3326_ _)) =>
- (match (string_drop _s3325_ _s3326_) with
- | _s3327_ =>
- (reg_name_matches_prefix _s3327_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5733_ _) =>
+ (match (string_drop _s5732_ _s5733_) with
+ | _s5734_ =>
+ (reg_name_matches_prefix _s5734_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s3328_ _)) =>
- (match (string_drop _s3327_ _s3328_) with
- | _s3329_ =>
- (sep_matches_prefix _s3329_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5735_ _) =>
+ (match (string_drop _s5734_ _s5735_) with
+ | _s5736_ =>
+ (sep_matches_prefix _s5736_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s3330_ _)) =>
- (match (string_drop _s3329_ _s3330_) with
- | _s3331_ =>
- (reg_name_matches_prefix _s3331_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5737_ _) =>
+ (match (string_drop _s5736_ _s5737_) with
+ | _s5738_ =>
+ (reg_name_matches_prefix _s5738_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s3332_ _)) =>
- (match (string_drop _s3331_ _s3332_) with
- | _s3333_ =>
- (sep_matches_prefix _s3333_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5739_ _) =>
+ (match (string_drop _s5738_ _s5739_) with
+ | _s5740_ =>
+ (sep_matches_prefix _s5740_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s3334_ _)) =>
- (match (string_drop _s3333_ _s3334_) with
- | _s3335_ =>
- (reg_name_matches_prefix _s3335_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5741_ _) =>
+ (match (string_drop _s5740_ _s5741_) with
+ | _s5742_ =>
+ (reg_name_matches_prefix _s5742_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs2, existT _ _s3336_ _)) =>
- match (string_drop _s3335_
- _s3336_) with
- | s_ =>
- Some
- ((s, rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s5743_ _) =>
+ match (string_drop _s5742_
+ _s5743_) with
+ | s_ =>
+ Some (s, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string)))
- else returnm (None : option ((bool * mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((bool * mword 5 * mword 5 * mword 5 * string))).
-Definition _s3303_ (_s3304_ : string)
+Definition _s5710_ (_s5711_ : string)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string))) :=
-
- (match _s3304_ with
- | _s3305_ =>
- (mul_mnemonic_matches_prefix _s3305_) >>= fun w__0 : option (((bool * bool * bool) * {n : Z & ArithFact (n >=
+ (match _s5711_ with
+ | _s5712_ =>
+ (mul_mnemonic_matches_prefix _s5712_) >>= fun w__0 : option (((bool * bool * bool) * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some (((high, signed1, signed2), existT _ _s3306_ _)) =>
- (match (string_drop _s3305_ _s3306_) with
- | _s3307_ =>
- (spc_matches_prefix _s3307_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some ((high, signed1, signed2), existT _ _s5713_ _) =>
+ (match (string_drop _s5712_ _s5713_) with
+ | _s5714_ =>
+ (spc_matches_prefix _s5714_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s3308_ _)) =>
- (match (string_drop _s3307_ _s3308_) with
- | _s3309_ =>
- (reg_name_matches_prefix _s3309_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5715_ _) =>
+ (match (string_drop _s5714_ _s5715_) with
+ | _s5716_ =>
+ (reg_name_matches_prefix _s5716_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s3310_ _)) =>
- (match (string_drop _s3309_ _s3310_) with
- | _s3311_ =>
- (sep_matches_prefix _s3311_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5717_ _) =>
+ (match (string_drop _s5716_ _s5717_) with
+ | _s5718_ =>
+ (sep_matches_prefix _s5718_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s3312_ _)) =>
- (match (string_drop _s3311_ _s3312_) with
- | _s3313_ =>
- (reg_name_matches_prefix _s3313_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5719_ _) =>
+ (match (string_drop _s5718_ _s5719_) with
+ | _s5720_ =>
+ (reg_name_matches_prefix _s5720_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s3314_ _)) =>
- (match (string_drop _s3313_ _s3314_) with
- | _s3315_ =>
- (sep_matches_prefix _s3315_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5721_ _) =>
+ (match (string_drop _s5720_ _s5721_) with
+ | _s5722_ =>
+ (sep_matches_prefix _s5722_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s3316_ _)) =>
- (match (string_drop _s3315_ _s3316_) with
- | _s3317_ =>
- (reg_name_matches_prefix _s3317_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5723_ _) =>
+ (match (string_drop _s5722_ _s5723_) with
+ | _s5724_ =>
+ (reg_name_matches_prefix _s5724_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s3318_ _)) =>
- match (string_drop _s3317_
- _s3318_) with
- | s_ =>
- Some
- ((high, signed1, signed2, rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s5725_ _) =>
+ match (string_drop _s5724_
+ _s5725_) with
+ | s_ =>
+ Some
+ (high, signed1, signed2, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string))).
-Definition _s3291_ (_s3292_ : string)
-: M (option ((mword 5 * mword 5 * string))) :=
-
- let _s3293_ := _s3292_ in
- (if ((string_startswith _s3293_ "c.add")) then
- (match (string_drop _s3293_ (projT1 (string_length "c.add"))) with
- | _s3294_ =>
- (spc_matches_prefix _s3294_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5698_ (_s5699_ : string) : M (option ((mword 5 * mword 5 * string))) :=
+ let _s5700_ := _s5699_ in
+ (if string_startswith _s5700_ "c.add" then
+ (match (string_drop _s5700_ (projT1 (string_length "c.add"))) with
+ | _s5701_ =>
+ (spc_matches_prefix _s5701_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3295_ _)) =>
- (match (string_drop _s3294_ _s3295_) with
- | _s3296_ =>
- (reg_name_matches_prefix _s3296_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5702_ _) =>
+ (match (string_drop _s5701_ _s5702_) with
+ | _s5703_ =>
+ (reg_name_matches_prefix _s5703_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3297_ _)) =>
- (match (string_drop _s3296_ _s3297_) with
- | _s3298_ =>
- (sep_matches_prefix _s3298_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5704_ _) =>
+ (match (string_drop _s5703_ _s5704_) with
+ | _s5705_ =>
+ (sep_matches_prefix _s5705_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3299_ _)) =>
- (match (string_drop _s3298_ _s3299_) with
- | _s3300_ =>
- (reg_name_matches_prefix _s3300_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5706_ _) =>
+ (match (string_drop _s5705_ _s5706_) with
+ | _s5707_ =>
+ (reg_name_matches_prefix _s5707_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s3301_ _)) =>
- match (string_drop _s3300_ _s3301_) with
- | s_ => Some ((rsd, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5708_ _) =>
+ match (string_drop _s5707_ _s5708_) with
+ | s_ => Some (rsd, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- else returnm (None : option ((mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * string))).
-Definition _s3287_ (_s3288_ : string)
-: option string :=
-
- let _s3289_ := _s3288_ in
- if ((string_startswith _s3289_ "c.ebreak")) then
- match (string_drop _s3289_ (projT1 (string_length "c.ebreak"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s3275_ (_s3276_ : string)
-: M (option ((mword 5 * mword 5 * string))) :=
-
- let _s3277_ := _s3276_ in
- (if ((string_startswith _s3277_ "c.mv")) then
- (match (string_drop _s3277_ (projT1 (string_length "c.mv"))) with
- | _s3278_ =>
- (spc_matches_prefix _s3278_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5694_ (_s5695_ : string) : option string :=
+ let _s5696_ := _s5695_ in
+ if string_startswith _s5696_ "c.ebreak" then
+ match (string_drop _s5696_ (projT1 (string_length "c.ebreak"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s5682_ (_s5683_ : string) : M (option ((mword 5 * mword 5 * string))) :=
+ let _s5684_ := _s5683_ in
+ (if string_startswith _s5684_ "c.mv" then
+ (match (string_drop _s5684_ (projT1 (string_length "c.mv"))) with
+ | _s5685_ =>
+ (spc_matches_prefix _s5685_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3279_ _)) =>
- (match (string_drop _s3278_ _s3279_) with
- | _s3280_ =>
- (reg_name_matches_prefix _s3280_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5686_ _) =>
+ (match (string_drop _s5685_ _s5686_) with
+ | _s5687_ =>
+ (reg_name_matches_prefix _s5687_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s3281_ _)) =>
- (match (string_drop _s3280_ _s3281_) with
- | _s3282_ =>
- (sep_matches_prefix _s3282_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5688_ _) =>
+ (match (string_drop _s5687_ _s5688_) with
+ | _s5689_ =>
+ (sep_matches_prefix _s5689_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3283_ _)) =>
- (match (string_drop _s3282_ _s3283_) with
- | _s3284_ =>
- (reg_name_matches_prefix _s3284_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5690_ _) =>
+ (match (string_drop _s5689_ _s5690_) with
+ | _s5691_ =>
+ (reg_name_matches_prefix _s5691_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s3285_ _)) =>
- match (string_drop _s3284_ _s3285_) with
- | s_ => Some ((rd, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5692_ _) =>
+ match (string_drop _s5691_ _s5692_) with
+ | s_ => Some (rd, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- else returnm (None : option ((mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * string))).
-Definition _s3267_ (_s3268_ : string)
-: M (option ((mword 5 * string))) :=
-
- let _s3269_ := _s3268_ in
- (if ((string_startswith _s3269_ "c.jalr")) then
- (match (string_drop _s3269_ (projT1 (string_length "c.jalr"))) with
- | _s3270_ =>
- (spc_matches_prefix _s3270_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5674_ (_s5675_ : string) : M (option ((mword 5 * string))) :=
+ let _s5676_ := _s5675_ in
+ (if string_startswith _s5676_ "c.jalr" then
+ (match (string_drop _s5676_ (projT1 (string_length "c.jalr"))) with
+ | _s5677_ =>
+ (spc_matches_prefix _s5677_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3271_ _)) =>
- (match (string_drop _s3270_ _s3271_) with
- | _s3272_ =>
- (reg_name_matches_prefix _s3272_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5678_ _) =>
+ (match (string_drop _s5677_ _s5678_) with
+ | _s5679_ =>
+ (reg_name_matches_prefix _s5679_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((rs1, existT _ _s3273_ _)) =>
- match (string_drop _s3272_ _s3273_) with | s_ => Some ((rs1, s_)) end
- | _ => None
- end)
- : option ((mword 5 * string)))
+ returnm (match w__1 with
+ | Some (rs1, existT _ _s5680_ _) =>
+ match (string_drop _s5679_ _s5680_) with | s_ => Some (rs1, s_) end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * string)))
end)
: M (option ((mword 5 * string)))
- else returnm (None : option ((mword 5 * string))))
+ else returnm None)
: M (option ((mword 5 * string))).
-Definition _s3259_ (_s3260_ : string)
-: M (option ((mword 5 * string))) :=
-
- let _s3261_ := _s3260_ in
- (if ((string_startswith _s3261_ "c.jr")) then
- (match (string_drop _s3261_ (projT1 (string_length "c.jr"))) with
- | _s3262_ =>
- (spc_matches_prefix _s3262_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5666_ (_s5667_ : string) : M (option ((mword 5 * string))) :=
+ let _s5668_ := _s5667_ in
+ (if string_startswith _s5668_ "c.jr" then
+ (match (string_drop _s5668_ (projT1 (string_length "c.jr"))) with
+ | _s5669_ =>
+ (spc_matches_prefix _s5669_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3263_ _)) =>
- (match (string_drop _s3262_ _s3263_) with
- | _s3264_ =>
- (reg_name_matches_prefix _s3264_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5670_ _) =>
+ (match (string_drop _s5669_ _s5670_) with
+ | _s5671_ =>
+ (reg_name_matches_prefix _s5671_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__1 with
- | Some ((rs1, existT _ _s3265_ _)) =>
- match (string_drop _s3264_ _s3265_) with | s_ => Some ((rs1, s_)) end
- | _ => None
- end)
- : option ((mword 5 * string)))
+ returnm (match w__1 with
+ | Some (rs1, existT _ _s5672_ _) =>
+ match (string_drop _s5671_ _s5672_) with | s_ => Some (rs1, s_) end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * string)))
end)
: M (option ((mword 5 * string)))
- else returnm (None : option ((mword 5 * string))))
+ else returnm None)
: M (option ((mword 5 * string))).
-Definition _s3247_ (_s3248_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s3249_ := _s3248_ in
- (if ((string_startswith _s3249_ "c.sdsp")) then
- (match (string_drop _s3249_ (projT1 (string_length "c.sdsp"))) with
- | _s3250_ =>
- (spc_matches_prefix _s3250_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5654_ (_s5655_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5656_ := _s5655_ in
+ (if string_startswith _s5656_ "c.sdsp" then
+ (match (string_drop _s5656_ (projT1 (string_length "c.sdsp"))) with
+ | _s5657_ =>
+ (spc_matches_prefix _s5657_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3251_ _)) =>
- (match (string_drop _s3250_ _s3251_) with
- | _s3252_ =>
- (reg_name_matches_prefix _s3252_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5658_ _) =>
+ (match (string_drop _s5657_ _s5658_) with
+ | _s5659_ =>
+ (reg_name_matches_prefix _s5659_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs2, existT _ _s3253_ _)) =>
- (match (string_drop _s3252_ _s3253_) with
- | _s3254_ =>
- (sep_matches_prefix _s3254_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s5660_ _) =>
+ (match (string_drop _s5659_ _s5660_) with
+ | _s5661_ =>
+ (sep_matches_prefix _s5661_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3255_ _)) =>
- match (string_drop _s3254_ _s3255_) with
- | _s3256_ =>
- match (hex_bits_6_matches_prefix _s3256_) with
- | Some ((uimm, existT _ _s3257_ _)) =>
- match (string_drop _s3256_ _s3257_) with
- | s_ => Some ((rs2, uimm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5662_ _) =>
+ match (string_drop _s5661_ _s5662_) with
+ | _s5663_ =>
+ match (hex_bits_6_matches_prefix _s5663_) with
+ | Some (uimm, existT _ _s5664_ _) =>
+ match (string_drop _s5663_ _s5664_) with
+ | s_ => Some (rs2, uimm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s3235_ (_s3236_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s3237_ := _s3236_ in
- (if ((string_startswith _s3237_ "c.swsp")) then
- (match (string_drop _s3237_ (projT1 (string_length "c.swsp"))) with
- | _s3238_ =>
- (spc_matches_prefix _s3238_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5642_ (_s5643_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5644_ := _s5643_ in
+ (if string_startswith _s5644_ "c.swsp" then
+ (match (string_drop _s5644_ (projT1 (string_length "c.swsp"))) with
+ | _s5645_ =>
+ (spc_matches_prefix _s5645_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3239_ _)) =>
- (match (string_drop _s3238_ _s3239_) with
- | _s3240_ =>
- (reg_name_matches_prefix _s3240_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5646_ _) =>
+ (match (string_drop _s5645_ _s5646_) with
+ | _s5647_ =>
+ (reg_name_matches_prefix _s5647_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s3241_ _)) =>
- (match (string_drop _s3240_ _s3241_) with
- | _s3242_ =>
- (sep_matches_prefix _s3242_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5648_ _) =>
+ (match (string_drop _s5647_ _s5648_) with
+ | _s5649_ =>
+ (sep_matches_prefix _s5649_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3243_ _)) =>
- match (string_drop _s3242_ _s3243_) with
- | _s3244_ =>
- match (hex_bits_6_matches_prefix _s3244_) with
- | Some ((uimm, existT _ _s3245_ _)) =>
- match (string_drop _s3244_ _s3245_) with
- | s_ => Some ((rd, uimm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5650_ _) =>
+ match (string_drop _s5649_ _s5650_) with
+ | _s5651_ =>
+ match (hex_bits_6_matches_prefix _s5651_) with
+ | Some (uimm, existT _ _s5652_ _) =>
+ match (string_drop _s5651_ _s5652_) with
+ | s_ => Some (rd, uimm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s3223_ (_s3224_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s3225_ := _s3224_ in
- (if ((string_startswith _s3225_ "c.ldsp")) then
- (match (string_drop _s3225_ (projT1 (string_length "c.ldsp"))) with
- | _s3226_ =>
- (spc_matches_prefix _s3226_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5630_ (_s5631_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5632_ := _s5631_ in
+ (if string_startswith _s5632_ "c.ldsp" then
+ (match (string_drop _s5632_ (projT1 (string_length "c.ldsp"))) with
+ | _s5633_ =>
+ (spc_matches_prefix _s5633_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3227_ _)) =>
- (match (string_drop _s3226_ _s3227_) with
- | _s3228_ =>
- (reg_name_matches_prefix _s3228_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5634_ _) =>
+ (match (string_drop _s5633_ _s5634_) with
+ | _s5635_ =>
+ (reg_name_matches_prefix _s5635_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s3229_ _)) =>
- (match (string_drop _s3228_ _s3229_) with
- | _s3230_ =>
- (sep_matches_prefix _s3230_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5636_ _) =>
+ (match (string_drop _s5635_ _s5636_) with
+ | _s5637_ =>
+ (sep_matches_prefix _s5637_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3231_ _)) =>
- match (string_drop _s3230_ _s3231_) with
- | _s3232_ =>
- match (hex_bits_6_matches_prefix _s3232_) with
- | Some ((uimm, existT _ _s3233_ _)) =>
- match (string_drop _s3232_ _s3233_) with
- | s_ => Some ((rd, uimm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5638_ _) =>
+ match (string_drop _s5637_ _s5638_) with
+ | _s5639_ =>
+ match (hex_bits_6_matches_prefix _s5639_) with
+ | Some (uimm, existT _ _s5640_ _) =>
+ match (string_drop _s5639_ _s5640_) with
+ | s_ => Some (rd, uimm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s3211_ (_s3212_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s3213_ := _s3212_ in
- (if ((string_startswith _s3213_ "c.lwsp")) then
- (match (string_drop _s3213_ (projT1 (string_length "c.lwsp"))) with
- | _s3214_ =>
- (spc_matches_prefix _s3214_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5618_ (_s5619_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5620_ := _s5619_ in
+ (if string_startswith _s5620_ "c.lwsp" then
+ (match (string_drop _s5620_ (projT1 (string_length "c.lwsp"))) with
+ | _s5621_ =>
+ (spc_matches_prefix _s5621_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3215_ _)) =>
- (match (string_drop _s3214_ _s3215_) with
- | _s3216_ =>
- (reg_name_matches_prefix _s3216_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5622_ _) =>
+ (match (string_drop _s5621_ _s5622_) with
+ | _s5623_ =>
+ (reg_name_matches_prefix _s5623_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s3217_ _)) =>
- (match (string_drop _s3216_ _s3217_) with
- | _s3218_ =>
- (sep_matches_prefix _s3218_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5624_ _) =>
+ (match (string_drop _s5623_ _s5624_) with
+ | _s5625_ =>
+ (sep_matches_prefix _s5625_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3219_ _)) =>
- match (string_drop _s3218_ _s3219_) with
- | _s3220_ =>
- match (hex_bits_6_matches_prefix _s3220_) with
- | Some ((uimm, existT _ _s3221_ _)) =>
- match (string_drop _s3220_ _s3221_) with
- | s_ => Some ((rd, uimm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5626_ _) =>
+ match (string_drop _s5625_ _s5626_) with
+ | _s5627_ =>
+ match (hex_bits_6_matches_prefix _s5627_) with
+ | Some (uimm, existT _ _s5628_ _) =>
+ match (string_drop _s5627_ _s5628_) with
+ | s_ => Some (rd, uimm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s3199_ (_s3200_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s3201_ := _s3200_ in
- (if ((string_startswith _s3201_ "c.slli")) then
- (match (string_drop _s3201_ (projT1 (string_length "c.slli"))) with
- | _s3202_ =>
- (spc_matches_prefix _s3202_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5606_ (_s5607_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5608_ := _s5607_ in
+ (if string_startswith _s5608_ "c.slli" then
+ (match (string_drop _s5608_ (projT1 (string_length "c.slli"))) with
+ | _s5609_ =>
+ (spc_matches_prefix _s5609_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3203_ _)) =>
- (match (string_drop _s3202_ _s3203_) with
- | _s3204_ =>
- (reg_name_matches_prefix _s3204_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5610_ _) =>
+ (match (string_drop _s5609_ _s5610_) with
+ | _s5611_ =>
+ (reg_name_matches_prefix _s5611_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3205_ _)) =>
- (match (string_drop _s3204_ _s3205_) with
- | _s3206_ =>
- (sep_matches_prefix _s3206_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5612_ _) =>
+ (match (string_drop _s5611_ _s5612_) with
+ | _s5613_ =>
+ (sep_matches_prefix _s5613_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3207_ _)) =>
- match (string_drop _s3206_ _s3207_) with
- | _s3208_ =>
- match (hex_bits_6_matches_prefix _s3208_) with
- | Some ((shamt, existT _ _s3209_ _)) =>
- match (string_drop _s3208_ _s3209_) with
- | s_ => Some ((rsd, shamt, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5614_ _) =>
+ match (string_drop _s5613_ _s5614_) with
+ | _s5615_ =>
+ match (hex_bits_6_matches_prefix _s5615_) with
+ | Some (shamt, existT _ _s5616_ _) =>
+ match (string_drop _s5615_ _s5616_) with
+ | s_ => Some (rsd, shamt, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s3187_ (_s3188_ : string)
-: M (option ((mword 3 * mword 8 * string))) :=
-
- let _s3189_ := _s3188_ in
- (if ((string_startswith _s3189_ "c.bnez")) then
- (match (string_drop _s3189_ (projT1 (string_length "c.bnez"))) with
- | _s3190_ =>
- (spc_matches_prefix _s3190_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5594_ (_s5595_ : string) : M (option ((mword 3 * mword 8 * string))) :=
+ let _s5596_ := _s5595_ in
+ (if string_startswith _s5596_ "c.bnez" then
+ (match (string_drop _s5596_ (projT1 (string_length "c.bnez"))) with
+ | _s5597_ =>
+ (spc_matches_prefix _s5597_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3191_ _)) =>
- (match (string_drop _s3190_ _s3191_) with
- | _s3192_ =>
- (creg_name_matches_prefix _s3192_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5598_ _) =>
+ (match (string_drop _s5597_ _s5598_) with
+ | _s5599_ =>
+ (creg_name_matches_prefix _s5599_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs, existT _ _s3193_ _)) =>
- (match (string_drop _s3192_ _s3193_) with
- | _s3194_ =>
- (sep_matches_prefix _s3194_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs, existT _ _s5600_ _) =>
+ (match (string_drop _s5599_ _s5600_) with
+ | _s5601_ =>
+ (sep_matches_prefix _s5601_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3195_ _)) =>
- match (string_drop _s3194_ _s3195_) with
- | _s3196_ =>
- match (hex_bits_8_matches_prefix _s3196_) with
- | Some ((imm, existT _ _s3197_ _)) =>
- match (string_drop _s3196_ _s3197_) with
- | s_ => Some ((rs, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5602_ _) =>
+ match (string_drop _s5601_ _s5602_) with
+ | _s5603_ =>
+ match (hex_bits_8_matches_prefix _s5603_) with
+ | Some (imm, existT _ _s5604_ _) =>
+ match (string_drop _s5603_ _s5604_) with
+ | s_ => Some (rs, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8 * string)))
- | _ => returnm (None : option ((mword 3 * mword 8 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8 * string)))
end)
: M (option ((mword 3 * mword 8 * string)))
- | _ => returnm (None : option ((mword 3 * mword 8 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8 * string)))
end)
: M (option ((mword 3 * mword 8 * string)))
- else returnm (None : option ((mword 3 * mword 8 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 8 * string))).
-Definition _s3175_ (_s3176_ : string)
-: M (option ((mword 3 * mword 8 * string))) :=
-
- let _s3177_ := _s3176_ in
- (if ((string_startswith _s3177_ "c.beqz")) then
- (match (string_drop _s3177_ (projT1 (string_length "c.beqz"))) with
- | _s3178_ =>
- (spc_matches_prefix _s3178_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5582_ (_s5583_ : string) : M (option ((mword 3 * mword 8 * string))) :=
+ let _s5584_ := _s5583_ in
+ (if string_startswith _s5584_ "c.beqz" then
+ (match (string_drop _s5584_ (projT1 (string_length "c.beqz"))) with
+ | _s5585_ =>
+ (spc_matches_prefix _s5585_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3179_ _)) =>
- (match (string_drop _s3178_ _s3179_) with
- | _s3180_ =>
- (creg_name_matches_prefix _s3180_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5586_ _) =>
+ (match (string_drop _s5585_ _s5586_) with
+ | _s5587_ =>
+ (creg_name_matches_prefix _s5587_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs, existT _ _s3181_ _)) =>
- (match (string_drop _s3180_ _s3181_) with
- | _s3182_ =>
- (sep_matches_prefix _s3182_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs, existT _ _s5588_ _) =>
+ (match (string_drop _s5587_ _s5588_) with
+ | _s5589_ =>
+ (sep_matches_prefix _s5589_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3183_ _)) =>
- match (string_drop _s3182_ _s3183_) with
- | _s3184_ =>
- match (hex_bits_8_matches_prefix _s3184_) with
- | Some ((imm, existT _ _s3185_ _)) =>
- match (string_drop _s3184_ _s3185_) with
- | s_ => Some ((rs, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5590_ _) =>
+ match (string_drop _s5589_ _s5590_) with
+ | _s5591_ =>
+ match (hex_bits_8_matches_prefix _s5591_) with
+ | Some (imm, existT _ _s5592_ _) =>
+ match (string_drop _s5591_ _s5592_) with
+ | s_ => Some (rs, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8 * string)))
- | _ => returnm (None : option ((mword 3 * mword 8 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8 * string)))
end)
: M (option ((mword 3 * mword 8 * string)))
- | _ => returnm (None : option ((mword 3 * mword 8 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8 * string)))
end)
: M (option ((mword 3 * mword 8 * string)))
- else returnm (None : option ((mword 3 * mword 8 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 8 * string))).
-Definition _s3167_ (_s3168_ : string)
-: M (option ((mword 11 * string))) :=
-
- let _s3169_ := _s3168_ in
- (if ((string_startswith _s3169_ "c.j")) then
- (match (string_drop _s3169_ (projT1 (string_length "c.j"))) with
- | _s3170_ =>
- (spc_matches_prefix _s3170_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s3171_ _)) =>
- match (string_drop _s3170_ _s3171_) with
- | _s3172_ =>
- match (hex_bits_11_matches_prefix _s3172_) with
- | Some ((imm, existT _ _s3173_ _)) =>
- match (string_drop _s3172_ _s3173_) with | s_ => Some ((imm, s_)) end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 11 * string)))
+Definition _s5574_ (_s5575_ : string) : M (option ((mword 11 * string))) :=
+ let _s5576_ := _s5575_ in
+ (if string_startswith _s5576_ "c.j" then
+ (match (string_drop _s5576_ (projT1 (string_length "c.j"))) with
+ | _s5577_ =>
+ (spc_matches_prefix _s5577_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s5578_ _) =>
+ match (string_drop _s5577_ _s5578_) with
+ | _s5579_ =>
+ match (hex_bits_11_matches_prefix _s5579_) with
+ | Some (imm, existT _ _s5580_ _) =>
+ match (string_drop _s5579_ _s5580_) with | s_ => Some (imm, s_) end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 11 * string)))
- else returnm (None : option ((mword 11 * string))))
+ else returnm None)
: M (option ((mword 11 * string))).
-Definition _s3155_ (_s3156_ : string)
-: M (option ((mword 3 * mword 3 * string))) :=
-
- let _s3157_ := _s3156_ in
- (if ((string_startswith _s3157_ "c.addw")) then
- (match (string_drop _s3157_ (projT1 (string_length "c.addw"))) with
- | _s3158_ =>
- (spc_matches_prefix _s3158_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5562_ (_s5563_ : string) : M (option ((mword 3 * mword 3 * string))) :=
+ let _s5564_ := _s5563_ in
+ (if string_startswith _s5564_ "c.addw" then
+ (match (string_drop _s5564_ (projT1 (string_length "c.addw"))) with
+ | _s5565_ =>
+ (spc_matches_prefix _s5565_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3159_ _)) =>
- (match (string_drop _s3158_ _s3159_) with
- | _s3160_ =>
- (creg_name_matches_prefix _s3160_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5566_ _) =>
+ (match (string_drop _s5565_ _s5566_) with
+ | _s5567_ =>
+ (creg_name_matches_prefix _s5567_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3161_ _)) =>
- (match (string_drop _s3160_ _s3161_) with
- | _s3162_ =>
- (sep_matches_prefix _s3162_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5568_ _) =>
+ (match (string_drop _s5567_ _s5568_) with
+ | _s5569_ =>
+ (sep_matches_prefix _s5569_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3163_ _)) =>
- (match (string_drop _s3162_ _s3163_) with
- | _s3164_ =>
- (creg_name_matches_prefix _s3164_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5570_ _) =>
+ (match (string_drop _s5569_ _s5570_) with
+ | _s5571_ =>
+ (creg_name_matches_prefix _s5571_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s3165_ _)) =>
- match (string_drop _s3164_ _s3165_) with
- | s_ => Some ((rsd, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5572_ _) =>
+ match (string_drop _s5571_ _s5572_) with
+ | s_ => Some (rsd, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- else returnm (None : option ((mword 3 * mword 3 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * string))).
-Definition _s3143_ (_s3144_ : string)
-: M (option ((mword 3 * mword 3 * string))) :=
-
- let _s3145_ := _s3144_ in
- (if ((string_startswith _s3145_ "c.subw")) then
- (match (string_drop _s3145_ (projT1 (string_length "c.subw"))) with
- | _s3146_ =>
- (spc_matches_prefix _s3146_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5550_ (_s5551_ : string) : M (option ((mword 3 * mword 3 * string))) :=
+ let _s5552_ := _s5551_ in
+ (if string_startswith _s5552_ "c.subw" then
+ (match (string_drop _s5552_ (projT1 (string_length "c.subw"))) with
+ | _s5553_ =>
+ (spc_matches_prefix _s5553_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3147_ _)) =>
- (match (string_drop _s3146_ _s3147_) with
- | _s3148_ =>
- (creg_name_matches_prefix _s3148_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5554_ _) =>
+ (match (string_drop _s5553_ _s5554_) with
+ | _s5555_ =>
+ (creg_name_matches_prefix _s5555_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3149_ _)) =>
- (match (string_drop _s3148_ _s3149_) with
- | _s3150_ =>
- (sep_matches_prefix _s3150_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5556_ _) =>
+ (match (string_drop _s5555_ _s5556_) with
+ | _s5557_ =>
+ (sep_matches_prefix _s5557_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3151_ _)) =>
- (match (string_drop _s3150_ _s3151_) with
- | _s3152_ =>
- (creg_name_matches_prefix _s3152_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5558_ _) =>
+ (match (string_drop _s5557_ _s5558_) with
+ | _s5559_ =>
+ (creg_name_matches_prefix _s5559_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s3153_ _)) =>
- match (string_drop _s3152_ _s3153_) with
- | s_ => Some ((rsd, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5560_ _) =>
+ match (string_drop _s5559_ _s5560_) with
+ | s_ => Some (rsd, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- else returnm (None : option ((mword 3 * mword 3 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * string))).
-Definition _s3131_ (_s3132_ : string)
-: M (option ((mword 3 * mword 3 * string))) :=
-
- let _s3133_ := _s3132_ in
- (if ((string_startswith _s3133_ "c.and")) then
- (match (string_drop _s3133_ (projT1 (string_length "c.and"))) with
- | _s3134_ =>
- (spc_matches_prefix _s3134_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5538_ (_s5539_ : string) : M (option ((mword 3 * mword 3 * string))) :=
+ let _s5540_ := _s5539_ in
+ (if string_startswith _s5540_ "c.and" then
+ (match (string_drop _s5540_ (projT1 (string_length "c.and"))) with
+ | _s5541_ =>
+ (spc_matches_prefix _s5541_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3135_ _)) =>
- (match (string_drop _s3134_ _s3135_) with
- | _s3136_ =>
- (creg_name_matches_prefix _s3136_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5542_ _) =>
+ (match (string_drop _s5541_ _s5542_) with
+ | _s5543_ =>
+ (creg_name_matches_prefix _s5543_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3137_ _)) =>
- (match (string_drop _s3136_ _s3137_) with
- | _s3138_ =>
- (sep_matches_prefix _s3138_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5544_ _) =>
+ (match (string_drop _s5543_ _s5544_) with
+ | _s5545_ =>
+ (sep_matches_prefix _s5545_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3139_ _)) =>
- (match (string_drop _s3138_ _s3139_) with
- | _s3140_ =>
- (creg_name_matches_prefix _s3140_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5546_ _) =>
+ (match (string_drop _s5545_ _s5546_) with
+ | _s5547_ =>
+ (creg_name_matches_prefix _s5547_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s3141_ _)) =>
- match (string_drop _s3140_ _s3141_) with
- | s_ => Some ((rsd, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5548_ _) =>
+ match (string_drop _s5547_ _s5548_) with
+ | s_ => Some (rsd, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- else returnm (None : option ((mword 3 * mword 3 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * string))).
-Definition _s3119_ (_s3120_ : string)
-: M (option ((mword 3 * mword 3 * string))) :=
-
- let _s3121_ := _s3120_ in
- (if ((string_startswith _s3121_ "c.or")) then
- (match (string_drop _s3121_ (projT1 (string_length "c.or"))) with
- | _s3122_ =>
- (spc_matches_prefix _s3122_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5526_ (_s5527_ : string) : M (option ((mword 3 * mword 3 * string))) :=
+ let _s5528_ := _s5527_ in
+ (if string_startswith _s5528_ "c.or" then
+ (match (string_drop _s5528_ (projT1 (string_length "c.or"))) with
+ | _s5529_ =>
+ (spc_matches_prefix _s5529_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3123_ _)) =>
- (match (string_drop _s3122_ _s3123_) with
- | _s3124_ =>
- (creg_name_matches_prefix _s3124_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5530_ _) =>
+ (match (string_drop _s5529_ _s5530_) with
+ | _s5531_ =>
+ (creg_name_matches_prefix _s5531_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3125_ _)) =>
- (match (string_drop _s3124_ _s3125_) with
- | _s3126_ =>
- (sep_matches_prefix _s3126_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5532_ _) =>
+ (match (string_drop _s5531_ _s5532_) with
+ | _s5533_ =>
+ (sep_matches_prefix _s5533_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3127_ _)) =>
- (match (string_drop _s3126_ _s3127_) with
- | _s3128_ =>
- (creg_name_matches_prefix _s3128_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5534_ _) =>
+ (match (string_drop _s5533_ _s5534_) with
+ | _s5535_ =>
+ (creg_name_matches_prefix _s5535_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s3129_ _)) =>
- match (string_drop _s3128_ _s3129_) with
- | s_ => Some ((rsd, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5536_ _) =>
+ match (string_drop _s5535_ _s5536_) with
+ | s_ => Some (rsd, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- else returnm (None : option ((mword 3 * mword 3 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * string))).
-Definition _s3107_ (_s3108_ : string)
-: M (option ((mword 3 * mword 3 * string))) :=
-
- let _s3109_ := _s3108_ in
- (if ((string_startswith _s3109_ "c.xor")) then
- (match (string_drop _s3109_ (projT1 (string_length "c.xor"))) with
- | _s3110_ =>
- (spc_matches_prefix _s3110_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5514_ (_s5515_ : string) : M (option ((mword 3 * mword 3 * string))) :=
+ let _s5516_ := _s5515_ in
+ (if string_startswith _s5516_ "c.xor" then
+ (match (string_drop _s5516_ (projT1 (string_length "c.xor"))) with
+ | _s5517_ =>
+ (spc_matches_prefix _s5517_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3111_ _)) =>
- (match (string_drop _s3110_ _s3111_) with
- | _s3112_ =>
- (creg_name_matches_prefix _s3112_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5518_ _) =>
+ (match (string_drop _s5517_ _s5518_) with
+ | _s5519_ =>
+ (creg_name_matches_prefix _s5519_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3113_ _)) =>
- (match (string_drop _s3112_ _s3113_) with
- | _s3114_ =>
- (sep_matches_prefix _s3114_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5520_ _) =>
+ (match (string_drop _s5519_ _s5520_) with
+ | _s5521_ =>
+ (sep_matches_prefix _s5521_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3115_ _)) =>
- (match (string_drop _s3114_ _s3115_) with
- | _s3116_ =>
- (creg_name_matches_prefix _s3116_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5522_ _) =>
+ (match (string_drop _s5521_ _s5522_) with
+ | _s5523_ =>
+ (creg_name_matches_prefix _s5523_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s3117_ _)) =>
- match (string_drop _s3116_ _s3117_) with
- | s_ => Some ((rsd, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5524_ _) =>
+ match (string_drop _s5523_ _s5524_) with
+ | s_ => Some (rsd, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- else returnm (None : option ((mword 3 * mword 3 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * string))).
-Definition _s3095_ (_s3096_ : string)
-: M (option ((mword 3 * mword 3 * string))) :=
-
- let _s3097_ := _s3096_ in
- (if ((string_startswith _s3097_ "c.sub")) then
- (match (string_drop _s3097_ (projT1 (string_length "c.sub"))) with
- | _s3098_ =>
- (spc_matches_prefix _s3098_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5502_ (_s5503_ : string) : M (option ((mword 3 * mword 3 * string))) :=
+ let _s5504_ := _s5503_ in
+ (if string_startswith _s5504_ "c.sub" then
+ (match (string_drop _s5504_ (projT1 (string_length "c.sub"))) with
+ | _s5505_ =>
+ (spc_matches_prefix _s5505_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3099_ _)) =>
- (match (string_drop _s3098_ _s3099_) with
- | _s3100_ =>
- (creg_name_matches_prefix _s3100_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5506_ _) =>
+ (match (string_drop _s5505_ _s5506_) with
+ | _s5507_ =>
+ (creg_name_matches_prefix _s5507_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3101_ _)) =>
- (match (string_drop _s3100_ _s3101_) with
- | _s3102_ =>
- (sep_matches_prefix _s3102_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5508_ _) =>
+ (match (string_drop _s5507_ _s5508_) with
+ | _s5509_ =>
+ (sep_matches_prefix _s5509_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s3103_ _)) =>
- (match (string_drop _s3102_ _s3103_) with
- | _s3104_ =>
- (creg_name_matches_prefix _s3104_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5510_ _) =>
+ (match (string_drop _s5509_ _s5510_) with
+ | _s5511_ =>
+ (creg_name_matches_prefix _s5511_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s3105_ _)) =>
- match (string_drop _s3104_ _s3105_) with
- | s_ => Some ((rsd, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5512_ _) =>
+ match (string_drop _s5511_ _s5512_) with
+ | s_ => Some (rsd, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * string)))
end)
: M (option ((mword 3 * mword 3 * string)))
- else returnm (None : option ((mword 3 * mword 3 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * string))).
-Definition _s3083_ (_s3084_ : string)
-: M (option ((mword 3 * mword 6 * string))) :=
-
- let _s3085_ := _s3084_ in
- (if ((string_startswith _s3085_ "c.andi")) then
- (match (string_drop _s3085_ (projT1 (string_length "c.andi"))) with
- | _s3086_ =>
- (spc_matches_prefix _s3086_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5490_ (_s5491_ : string) : M (option ((mword 3 * mword 6 * string))) :=
+ let _s5492_ := _s5491_ in
+ (if string_startswith _s5492_ "c.andi" then
+ (match (string_drop _s5492_ (projT1 (string_length "c.andi"))) with
+ | _s5493_ =>
+ (spc_matches_prefix _s5493_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3087_ _)) =>
- (match (string_drop _s3086_ _s3087_) with
- | _s3088_ =>
- (creg_name_matches_prefix _s3088_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5494_ _) =>
+ (match (string_drop _s5493_ _s5494_) with
+ | _s5495_ =>
+ (creg_name_matches_prefix _s5495_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3089_ _)) =>
- (match (string_drop _s3088_ _s3089_) with
- | _s3090_ =>
- (sep_matches_prefix _s3090_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5496_ _) =>
+ (match (string_drop _s5495_ _s5496_) with
+ | _s5497_ =>
+ (sep_matches_prefix _s5497_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3091_ _)) =>
- match (string_drop _s3090_ _s3091_) with
- | _s3092_ =>
- match (hex_bits_6_matches_prefix _s3092_) with
- | Some ((imm, existT _ _s3093_ _)) =>
- match (string_drop _s3092_ _s3093_) with
- | s_ => Some ((rsd, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5498_ _) =>
+ match (string_drop _s5497_ _s5498_) with
+ | _s5499_ =>
+ match (hex_bits_6_matches_prefix _s5499_) with
+ | Some (imm, existT _ _s5500_ _) =>
+ match (string_drop _s5499_ _s5500_) with
+ | s_ => Some (rsd, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6 * string)))
- | _ => returnm (None : option ((mword 3 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6 * string)))
end)
: M (option ((mword 3 * mword 6 * string)))
- | _ => returnm (None : option ((mword 3 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6 * string)))
end)
: M (option ((mword 3 * mword 6 * string)))
- else returnm (None : option ((mword 3 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 6 * string))).
-Definition _s3071_ (_s3072_ : string)
-: M (option ((mword 3 * mword 6 * string))) :=
-
- let _s3073_ := _s3072_ in
- (if ((string_startswith _s3073_ "c.srai")) then
- (match (string_drop _s3073_ (projT1 (string_length "c.srai"))) with
- | _s3074_ =>
- (spc_matches_prefix _s3074_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5478_ (_s5479_ : string) : M (option ((mword 3 * mword 6 * string))) :=
+ let _s5480_ := _s5479_ in
+ (if string_startswith _s5480_ "c.srai" then
+ (match (string_drop _s5480_ (projT1 (string_length "c.srai"))) with
+ | _s5481_ =>
+ (spc_matches_prefix _s5481_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3075_ _)) =>
- (match (string_drop _s3074_ _s3075_) with
- | _s3076_ =>
- (creg_name_matches_prefix _s3076_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5482_ _) =>
+ (match (string_drop _s5481_ _s5482_) with
+ | _s5483_ =>
+ (creg_name_matches_prefix _s5483_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3077_ _)) =>
- (match (string_drop _s3076_ _s3077_) with
- | _s3078_ =>
- (sep_matches_prefix _s3078_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5484_ _) =>
+ (match (string_drop _s5483_ _s5484_) with
+ | _s5485_ =>
+ (sep_matches_prefix _s5485_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3079_ _)) =>
- match (string_drop _s3078_ _s3079_) with
- | _s3080_ =>
- match (hex_bits_6_matches_prefix _s3080_) with
- | Some ((shamt, existT _ _s3081_ _)) =>
- match (string_drop _s3080_ _s3081_) with
- | s_ => Some ((rsd, shamt, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5486_ _) =>
+ match (string_drop _s5485_ _s5486_) with
+ | _s5487_ =>
+ match (hex_bits_6_matches_prefix _s5487_) with
+ | Some (shamt, existT _ _s5488_ _) =>
+ match (string_drop _s5487_ _s5488_) with
+ | s_ => Some (rsd, shamt, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6 * string)))
- | _ => returnm (None : option ((mword 3 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6 * string)))
end)
: M (option ((mword 3 * mword 6 * string)))
- | _ => returnm (None : option ((mword 3 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6 * string)))
end)
: M (option ((mword 3 * mword 6 * string)))
- else returnm (None : option ((mword 3 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 6 * string))).
-Definition _s3059_ (_s3060_ : string)
-: M (option ((mword 3 * mword 6 * string))) :=
-
- let _s3061_ := _s3060_ in
- (if ((string_startswith _s3061_ "c.srli")) then
- (match (string_drop _s3061_ (projT1 (string_length "c.srli"))) with
- | _s3062_ =>
- (spc_matches_prefix _s3062_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5466_ (_s5467_ : string) : M (option ((mword 3 * mword 6 * string))) :=
+ let _s5468_ := _s5467_ in
+ (if string_startswith _s5468_ "c.srli" then
+ (match (string_drop _s5468_ (projT1 (string_length "c.srli"))) with
+ | _s5469_ =>
+ (spc_matches_prefix _s5469_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3063_ _)) =>
- (match (string_drop _s3062_ _s3063_) with
- | _s3064_ =>
- (creg_name_matches_prefix _s3064_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5470_ _) =>
+ (match (string_drop _s5469_ _s5470_) with
+ | _s5471_ =>
+ (creg_name_matches_prefix _s5471_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3065_ _)) =>
- (match (string_drop _s3064_ _s3065_) with
- | _s3066_ =>
- (sep_matches_prefix _s3066_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5472_ _) =>
+ (match (string_drop _s5471_ _s5472_) with
+ | _s5473_ =>
+ (sep_matches_prefix _s5473_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3067_ _)) =>
- match (string_drop _s3066_ _s3067_) with
- | _s3068_ =>
- match (hex_bits_6_matches_prefix _s3068_) with
- | Some ((shamt, existT _ _s3069_ _)) =>
- match (string_drop _s3068_ _s3069_) with
- | s_ => Some ((rsd, shamt, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5474_ _) =>
+ match (string_drop _s5473_ _s5474_) with
+ | _s5475_ =>
+ match (hex_bits_6_matches_prefix _s5475_) with
+ | Some (shamt, existT _ _s5476_ _) =>
+ match (string_drop _s5475_ _s5476_) with
+ | s_ => Some (rsd, shamt, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 6 * string)))
- | _ => returnm (None : option ((mword 3 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6 * string)))
end)
: M (option ((mword 3 * mword 6 * string)))
- | _ => returnm (None : option ((mword 3 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 6 * string)))
end)
: M (option ((mword 3 * mword 6 * string)))
- else returnm (None : option ((mword 3 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 6 * string))).
-Definition _s3047_ (_s3048_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s3049_ := _s3048_ in
- (if ((string_startswith _s3049_ "c.lui")) then
- (match (string_drop _s3049_ (projT1 (string_length "c.lui"))) with
- | _s3050_ =>
- (spc_matches_prefix _s3050_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5454_ (_s5455_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5456_ := _s5455_ in
+ (if string_startswith _s5456_ "c.lui" then
+ (match (string_drop _s5456_ (projT1 (string_length "c.lui"))) with
+ | _s5457_ =>
+ (spc_matches_prefix _s5457_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3051_ _)) =>
- (match (string_drop _s3050_ _s3051_) with
- | _s3052_ =>
- (reg_name_matches_prefix _s3052_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5458_ _) =>
+ (match (string_drop _s5457_ _s5458_) with
+ | _s5459_ =>
+ (reg_name_matches_prefix _s5459_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s3053_ _)) =>
- (match (string_drop _s3052_ _s3053_) with
- | _s3054_ =>
- (sep_matches_prefix _s3054_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5460_ _) =>
+ (match (string_drop _s5459_ _s5460_) with
+ | _s5461_ =>
+ (sep_matches_prefix _s5461_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3055_ _)) =>
- match (string_drop _s3054_ _s3055_) with
- | _s3056_ =>
- match (hex_bits_6_matches_prefix _s3056_) with
- | Some ((imm, existT _ _s3057_ _)) =>
- match (string_drop _s3056_ _s3057_) with
- | s_ => Some ((rd, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5462_ _) =>
+ match (string_drop _s5461_ _s5462_) with
+ | _s5463_ =>
+ match (hex_bits_6_matches_prefix _s5463_) with
+ | Some (imm, existT _ _s5464_ _) =>
+ match (string_drop _s5463_ _s5464_) with
+ | s_ => Some (rd, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s3039_ (_s3040_ : string)
-: M (option ((mword 6 * string))) :=
-
- let _s3041_ := _s3040_ in
- (if ((string_startswith _s3041_ "c.addi16sp")) then
- (match (string_drop _s3041_ (projT1 (string_length "c.addi16sp"))) with
- | _s3042_ =>
- (spc_matches_prefix _s3042_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s3043_ _)) =>
- match (string_drop _s3042_ _s3043_) with
- | _s3044_ =>
- match (hex_bits_6_matches_prefix _s3044_) with
- | Some ((imm, existT _ _s3045_ _)) =>
- match (string_drop _s3044_ _s3045_) with | s_ => Some ((imm, s_)) end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 6 * string)))
+Definition _s5446_ (_s5447_ : string) : M (option ((mword 6 * string))) :=
+ let _s5448_ := _s5447_ in
+ (if string_startswith _s5448_ "c.addi16sp" then
+ (match (string_drop _s5448_ (projT1 (string_length "c.addi16sp"))) with
+ | _s5449_ =>
+ (spc_matches_prefix _s5449_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s5450_ _) =>
+ match (string_drop _s5449_ _s5450_) with
+ | _s5451_ =>
+ match (hex_bits_6_matches_prefix _s5451_) with
+ | Some (imm, existT _ _s5452_ _) =>
+ match (string_drop _s5451_ _s5452_) with | s_ => Some (imm, s_) end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 6 * string)))
- else returnm (None : option ((mword 6 * string))))
+ else returnm None)
: M (option ((mword 6 * string))).
-Definition _s3027_ (_s3028_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s3029_ := _s3028_ in
- (if ((string_startswith _s3029_ "c.li")) then
- (match (string_drop _s3029_ (projT1 (string_length "c.li"))) with
- | _s3030_ =>
- (spc_matches_prefix _s3030_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5434_ (_s5435_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5436_ := _s5435_ in
+ (if string_startswith _s5436_ "c.li" then
+ (match (string_drop _s5436_ (projT1 (string_length "c.li"))) with
+ | _s5437_ =>
+ (spc_matches_prefix _s5437_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3031_ _)) =>
- (match (string_drop _s3030_ _s3031_) with
- | _s3032_ =>
- (reg_name_matches_prefix _s3032_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5438_ _) =>
+ (match (string_drop _s5437_ _s5438_) with
+ | _s5439_ =>
+ (reg_name_matches_prefix _s5439_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s3033_ _)) =>
- (match (string_drop _s3032_ _s3033_) with
- | _s3034_ =>
- (sep_matches_prefix _s3034_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5440_ _) =>
+ (match (string_drop _s5439_ _s5440_) with
+ | _s5441_ =>
+ (sep_matches_prefix _s5441_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3035_ _)) =>
- match (string_drop _s3034_ _s3035_) with
- | _s3036_ =>
- match (hex_bits_6_matches_prefix _s3036_) with
- | Some ((imm, existT _ _s3037_ _)) =>
- match (string_drop _s3036_ _s3037_) with
- | s_ => Some ((rd, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5442_ _) =>
+ match (string_drop _s5441_ _s5442_) with
+ | _s5443_ =>
+ match (hex_bits_6_matches_prefix _s5443_) with
+ | Some (imm, existT _ _s5444_ _) =>
+ match (string_drop _s5443_ _s5444_) with
+ | s_ => Some (rd, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s3015_ (_s3016_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s3017_ := _s3016_ in
- (if ((string_startswith _s3017_ "c.addiw")) then
- (match (string_drop _s3017_ (projT1 (string_length "c.addiw"))) with
- | _s3018_ =>
- (spc_matches_prefix _s3018_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5422_ (_s5423_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5424_ := _s5423_ in
+ (if string_startswith _s5424_ "c.addiw" then
+ (match (string_drop _s5424_ (projT1 (string_length "c.addiw"))) with
+ | _s5425_ =>
+ (spc_matches_prefix _s5425_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s3019_ _)) =>
- (match (string_drop _s3018_ _s3019_) with
- | _s3020_ =>
- (reg_name_matches_prefix _s3020_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5426_ _) =>
+ (match (string_drop _s5425_ _s5426_) with
+ | _s5427_ =>
+ (reg_name_matches_prefix _s5427_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3021_ _)) =>
- (match (string_drop _s3020_ _s3021_) with
- | _s3022_ =>
- (sep_matches_prefix _s3022_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5428_ _) =>
+ (match (string_drop _s5427_ _s5428_) with
+ | _s5429_ =>
+ (sep_matches_prefix _s5429_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3023_ _)) =>
- match (string_drop _s3022_ _s3023_) with
- | _s3024_ =>
- match (hex_bits_6_matches_prefix _s3024_) with
- | Some ((imm, existT _ _s3025_ _)) =>
- match (string_drop _s3024_ _s3025_) with
- | s_ => Some ((rsd, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5430_ _) =>
+ match (string_drop _s5429_ _s5430_) with
+ | _s5431_ =>
+ match (hex_bits_6_matches_prefix _s5431_) with
+ | Some (imm, existT _ _s5432_ _) =>
+ match (string_drop _s5431_ _s5432_) with
+ | s_ => Some (rsd, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s3007_ (_s3008_ : string)
-: M (option ((mword 11 * string))) :=
-
- let _s3009_ := _s3008_ in
- (if ((string_startswith _s3009_ "c.jal")) then
- (match (string_drop _s3009_ (projT1 (string_length "c.jal"))) with
- | _s3010_ =>
- (spc_matches_prefix _s3010_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
- returnm ((match w__0 with
- | Some ((tt, existT _ _s3011_ _)) =>
- match (string_drop _s3010_ _s3011_) with
- | _s3012_ =>
- match (hex_bits_12_matches_prefix _s3012_) with
- | Some ((v__826, existT _ _s3013_ _)) =>
- if ((eq_vec (subrange_vec_dec v__826 0 0)
- (vec_of_bits [B0] : mword (0 - 0 + 1)))) then
- let imm : mword 11 := subrange_vec_dec v__826 11 1 in
- let imm : mword 11 := subrange_vec_dec v__826 11 1 in
- match (string_drop _s3012_ _s3013_) with | s_ => Some ((imm, s_)) end
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 11 * string)))
+Definition _s5414_ (_s5415_ : string) : M (option ((mword 11 * string))) :=
+ let _s5416_ := _s5415_ in
+ (if string_startswith _s5416_ "c.jal" then
+ (match (string_drop _s5416_ (projT1 (string_length "c.jal"))) with
+ | _s5417_ =>
+ (spc_matches_prefix _s5417_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
+ returnm (match w__0 with
+ | Some (tt, existT _ _s5418_ _) =>
+ match (string_drop _s5417_ _s5418_) with
+ | _s5419_ =>
+ match (hex_bits_12_matches_prefix _s5419_) with
+ | Some (v__1386, existT _ _s5420_ _) =>
+ if eq_vec (subrange_vec_dec v__1386 0 0) ('b"0" : mword (0 - 0 + 1))
+ then
+ let imm : mword 11 := subrange_vec_dec v__1386 11 1 in
+ let imm : mword 11 := subrange_vec_dec v__1386 11 1 in
+ match (string_drop _s5419_ _s5420_) with | s_ => Some (imm, s_) end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 11 * string)))
- else returnm (None : option ((mword 11 * string))))
+ else returnm None)
: M (option ((mword 11 * string))).
-Definition _s2995_ (_s2996_ : string)
-: M (option ((mword 5 * mword 6 * string))) :=
-
- let _s2997_ := _s2996_ in
- (if ((string_startswith _s2997_ "c.addi")) then
- (match (string_drop _s2997_ (projT1 (string_length "c.addi"))) with
- | _s2998_ =>
- (spc_matches_prefix _s2998_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5402_ (_s5403_ : string) : M (option ((mword 5 * mword 6 * string))) :=
+ let _s5404_ := _s5403_ in
+ (if string_startswith _s5404_ "c.addi" then
+ (match (string_drop _s5404_ (projT1 (string_length "c.addi"))) with
+ | _s5405_ =>
+ (spc_matches_prefix _s5405_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2999_ _)) =>
- (match (string_drop _s2998_ _s2999_) with
- | _s3000_ =>
- (reg_name_matches_prefix _s3000_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5406_ _) =>
+ (match (string_drop _s5405_ _s5406_) with
+ | _s5407_ =>
+ (reg_name_matches_prefix _s5407_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsd, existT _ _s3001_ _)) =>
- (match (string_drop _s3000_ _s3001_) with
- | _s3002_ =>
- (sep_matches_prefix _s3002_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsd, existT _ _s5408_ _) =>
+ (match (string_drop _s5407_ _s5408_) with
+ | _s5409_ =>
+ (sep_matches_prefix _s5409_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s3003_ _)) =>
- match (string_drop _s3002_ _s3003_) with
- | _s3004_ =>
- match (hex_bits_6_matches_prefix _s3004_) with
- | Some ((nzi, existT _ _s3005_ _)) =>
- match (string_drop _s3004_ _s3005_) with
- | s_ => Some ((rsd, nzi, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 6 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5410_ _) =>
+ match (string_drop _s5409_ _s5410_) with
+ | _s5411_ =>
+ match (hex_bits_6_matches_prefix _s5411_) with
+ | Some (nzi, existT _ _s5412_ _) =>
+ match (string_drop _s5411_ _s5412_) with
+ | s_ => Some (rsd, nzi, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 6 * string)))
end)
: M (option ((mword 5 * mword 6 * string)))
- else returnm (None : option ((mword 5 * mword 6 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 6 * string))).
-Definition _s2979_ (_s2980_ : string)
-: M (option ((mword 3 * mword 3 * mword 5 * string))) :=
-
- let _s2981_ := _s2980_ in
- (if ((string_startswith _s2981_ "c.sd")) then
- (match (string_drop _s2981_ (projT1 (string_length "c.sd"))) with
- | _s2982_ =>
- (spc_matches_prefix _s2982_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5386_ (_s5387_ : string) : M (option ((mword 3 * mword 3 * mword 5 * string))) :=
+ let _s5388_ := _s5387_ in
+ (if string_startswith _s5388_ "c.sd" then
+ (match (string_drop _s5388_ (projT1 (string_length "c.sd"))) with
+ | _s5389_ =>
+ (spc_matches_prefix _s5389_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2983_ _)) =>
- (match (string_drop _s2982_ _s2983_) with
- | _s2984_ =>
- (creg_name_matches_prefix _s2984_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5390_ _) =>
+ (match (string_drop _s5389_ _s5390_) with
+ | _s5391_ =>
+ (creg_name_matches_prefix _s5391_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsc1, existT _ _s2985_ _)) =>
- (match (string_drop _s2984_ _s2985_) with
- | _s2986_ =>
- (sep_matches_prefix _s2986_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc1, existT _ _s5392_ _) =>
+ (match (string_drop _s5391_ _s5392_) with
+ | _s5393_ =>
+ (sep_matches_prefix _s5393_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2987_ _)) =>
- (match (string_drop _s2986_ _s2987_) with
- | _s2988_ =>
- (creg_name_matches_prefix _s2988_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5394_ _) =>
+ (match (string_drop _s5393_ _s5394_) with
+ | _s5395_ =>
+ (creg_name_matches_prefix _s5395_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc2, existT _ _s2989_ _)) =>
- (match (string_drop _s2988_ _s2989_) with
- | _s2990_ =>
- (sep_matches_prefix _s2990_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc2, existT _ _s5396_ _) =>
+ (match (string_drop _s5395_ _s5396_) with
+ | _s5397_ =>
+ (sep_matches_prefix _s5397_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2991_ _)) =>
- match (string_drop _s2990_ _s2991_) with
- | _s2992_ =>
- match (hex_bits_8_matches_prefix _s2992_) with
- | Some ((v__828, existT _ _s2993_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__828 2 0)
- (vec_of_bits [B0;B0;B0]
- : mword (2 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__828 7 3 in
- let uimm : mword 5 :=
- subrange_vec_dec v__828 7 3 in
- match (string_drop _s2992_ _s2993_) with
- | s_ => Some ((rsc1, rsc2, uimm, s_))
- end
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5 * string)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s5398_ _) =>
+ match (string_drop _s5397_ _s5398_) with
+ | _s5399_ =>
+ match (hex_bits_8_matches_prefix _s5399_) with
+ | Some (v__1388, existT _ _s5400_ _) =>
+ if eq_vec (subrange_vec_dec v__1388 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1388 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1388 7 3 in
+ match (string_drop _s5399_ _s5400_) with
+ | s_ => Some (rsc1, rsc2, uimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5 * string))).
-Definition _s2963_ (_s2964_ : string)
-: M (option ((mword 3 * mword 3 * mword 5 * string))) :=
-
- let _s2965_ := _s2964_ in
- (if ((string_startswith _s2965_ "c.sw")) then
- (match (string_drop _s2965_ (projT1 (string_length "c.sw"))) with
- | _s2966_ =>
- (spc_matches_prefix _s2966_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5370_ (_s5371_ : string) : M (option ((mword 3 * mword 3 * mword 5 * string))) :=
+ let _s5372_ := _s5371_ in
+ (if string_startswith _s5372_ "c.sw" then
+ (match (string_drop _s5372_ (projT1 (string_length "c.sw"))) with
+ | _s5373_ =>
+ (spc_matches_prefix _s5373_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2967_ _)) =>
- (match (string_drop _s2966_ _s2967_) with
- | _s2968_ =>
- (creg_name_matches_prefix _s2968_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5374_ _) =>
+ (match (string_drop _s5373_ _s5374_) with
+ | _s5375_ =>
+ (creg_name_matches_prefix _s5375_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rsc1, existT _ _s2969_ _)) =>
- (match (string_drop _s2968_ _s2969_) with
- | _s2970_ =>
- (sep_matches_prefix _s2970_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc1, existT _ _s5376_ _) =>
+ (match (string_drop _s5375_ _s5376_) with
+ | _s5377_ =>
+ (sep_matches_prefix _s5377_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2971_ _)) =>
- (match (string_drop _s2970_ _s2971_) with
- | _s2972_ =>
- (creg_name_matches_prefix _s2972_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5378_ _) =>
+ (match (string_drop _s5377_ _s5378_) with
+ | _s5379_ =>
+ (creg_name_matches_prefix _s5379_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc2, existT _ _s2973_ _)) =>
- (match (string_drop _s2972_ _s2973_) with
- | _s2974_ =>
- (sep_matches_prefix _s2974_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc2, existT _ _s5380_ _) =>
+ (match (string_drop _s5379_ _s5380_) with
+ | _s5381_ =>
+ (sep_matches_prefix _s5381_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2975_ _)) =>
- match (string_drop _s2974_ _s2975_) with
- | _s2976_ =>
- match (hex_bits_7_matches_prefix _s2976_) with
- | Some ((v__830, existT _ _s2977_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__830 1 0)
- (vec_of_bits [B0;B0]
- : mword (1 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__830 6 2 in
- let uimm : mword 5 :=
- subrange_vec_dec v__830 6 2 in
- match (string_drop _s2976_ _s2977_) with
- | s_ => Some ((rsc1, rsc2, uimm, s_))
- end
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5 * string)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s5382_ _) =>
+ match (string_drop _s5381_ _s5382_) with
+ | _s5383_ =>
+ match (hex_bits_7_matches_prefix _s5383_) with
+ | Some (v__1390, existT _ _s5384_ _) =>
+ if eq_vec (subrange_vec_dec v__1390 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1390 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1390 6 2 in
+ match (string_drop _s5383_ _s5384_) with
+ | s_ => Some (rsc1, rsc2, uimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5 * string))).
-Definition _s2947_ (_s2948_ : string)
-: M (option ((mword 3 * mword 3 * mword 5 * string))) :=
-
- let _s2949_ := _s2948_ in
- (if ((string_startswith _s2949_ "c.ld")) then
- (match (string_drop _s2949_ (projT1 (string_length "c.ld"))) with
- | _s2950_ =>
- (spc_matches_prefix _s2950_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5354_ (_s5355_ : string) : M (option ((mword 3 * mword 3 * mword 5 * string))) :=
+ let _s5356_ := _s5355_ in
+ (if string_startswith _s5356_ "c.ld" then
+ (match (string_drop _s5356_ (projT1 (string_length "c.ld"))) with
+ | _s5357_ =>
+ (spc_matches_prefix _s5357_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2951_ _)) =>
- (match (string_drop _s2950_ _s2951_) with
- | _s2952_ =>
- (creg_name_matches_prefix _s2952_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5358_ _) =>
+ (match (string_drop _s5357_ _s5358_) with
+ | _s5359_ =>
+ (creg_name_matches_prefix _s5359_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s2953_ _)) =>
- (match (string_drop _s2952_ _s2953_) with
- | _s2954_ =>
- (sep_matches_prefix _s2954_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s5360_ _) =>
+ (match (string_drop _s5359_ _s5360_) with
+ | _s5361_ =>
+ (sep_matches_prefix _s5361_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2955_ _)) =>
- (match (string_drop _s2954_ _s2955_) with
- | _s2956_ =>
- (creg_name_matches_prefix _s2956_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5362_ _) =>
+ (match (string_drop _s5361_ _s5362_) with
+ | _s5363_ =>
+ (creg_name_matches_prefix _s5363_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc, existT _ _s2957_ _)) =>
- (match (string_drop _s2956_ _s2957_) with
- | _s2958_ =>
- (sep_matches_prefix _s2958_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc, existT _ _s5364_ _) =>
+ (match (string_drop _s5363_ _s5364_) with
+ | _s5365_ =>
+ (sep_matches_prefix _s5365_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2959_ _)) =>
- match (string_drop _s2958_ _s2959_) with
- | _s2960_ =>
- match (hex_bits_8_matches_prefix _s2960_) with
- | Some ((v__832, existT _ _s2961_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__832 2 0)
- (vec_of_bits [B0;B0;B0]
- : mword (2 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__832 7 3 in
- let uimm : mword 5 :=
- subrange_vec_dec v__832 7 3 in
- match (string_drop _s2960_ _s2961_) with
- | s_ => Some ((rdc, rsc, uimm, s_))
- end
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5 * string)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s5366_ _) =>
+ match (string_drop _s5365_ _s5366_) with
+ | _s5367_ =>
+ match (hex_bits_8_matches_prefix _s5367_) with
+ | Some (v__1392, existT _ _s5368_ _) =>
+ if eq_vec (subrange_vec_dec v__1392 2 0)
+ ('b"000"
+ : mword (2 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1392 7 3 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1392 7 3 in
+ match (string_drop _s5367_ _s5368_) with
+ | s_ => Some (rdc, rsc, uimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5 * string))).
-Definition _s2931_ (_s2932_ : string)
-: M (option ((mword 3 * mword 3 * mword 5 * string))) :=
-
- let _s2933_ := _s2932_ in
- (if ((string_startswith _s2933_ "c.lw")) then
- (match (string_drop _s2933_ (projT1 (string_length "c.lw"))) with
- | _s2934_ =>
- (spc_matches_prefix _s2934_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5338_ (_s5339_ : string) : M (option ((mword 3 * mword 3 * mword 5 * string))) :=
+ let _s5340_ := _s5339_ in
+ (if string_startswith _s5340_ "c.lw" then
+ (match (string_drop _s5340_ (projT1 (string_length "c.lw"))) with
+ | _s5341_ =>
+ (spc_matches_prefix _s5341_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2935_ _)) =>
- (match (string_drop _s2934_ _s2935_) with
- | _s2936_ =>
- (creg_name_matches_prefix _s2936_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5342_ _) =>
+ (match (string_drop _s5341_ _s5342_) with
+ | _s5343_ =>
+ (creg_name_matches_prefix _s5343_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s2937_ _)) =>
- (match (string_drop _s2936_ _s2937_) with
- | _s2938_ =>
- (sep_matches_prefix _s2938_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s5344_ _) =>
+ (match (string_drop _s5343_ _s5344_) with
+ | _s5345_ =>
+ (sep_matches_prefix _s5345_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2939_ _)) =>
- (match (string_drop _s2938_ _s2939_) with
- | _s2940_ =>
- (creg_name_matches_prefix _s2940_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5346_ _) =>
+ (match (string_drop _s5345_ _s5346_) with
+ | _s5347_ =>
+ (creg_name_matches_prefix _s5347_) >>= fun w__3 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rsc, existT _ _s2941_ _)) =>
- (match (string_drop _s2940_ _s2941_) with
- | _s2942_ =>
- (sep_matches_prefix _s2942_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rsc, existT _ _s5348_ _) =>
+ (match (string_drop _s5347_ _s5348_) with
+ | _s5349_ =>
+ (sep_matches_prefix _s5349_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2943_ _)) =>
- match (string_drop _s2942_ _s2943_) with
- | _s2944_ =>
- match (hex_bits_7_matches_prefix _s2944_) with
- | Some ((v__834, existT _ _s2945_ _)) =>
- if ((eq_vec
- (subrange_vec_dec v__834 1 0)
- (vec_of_bits [B0;B0]
- : mword (1 - 0 + 1)))) then
- let uimm : mword 5 :=
- subrange_vec_dec v__834 6 2 in
- let uimm : mword 5 :=
- subrange_vec_dec v__834 6 2 in
- match (string_drop _s2944_ _s2945_) with
- | s_ => Some ((rdc, rsc, uimm, s_))
- end
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 3 * mword 5 * string)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s5350_ _) =>
+ match (string_drop _s5349_ _s5350_) with
+ | _s5351_ =>
+ match (hex_bits_7_matches_prefix _s5351_) with
+ | Some (v__1394, existT _ _s5352_ _) =>
+ if eq_vec (subrange_vec_dec v__1394 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1394 6 2 in
+ let uimm : mword 5 :=
+ subrange_vec_dec v__1394 6 2 in
+ match (string_drop _s5351_ _s5352_) with
+ | s_ => Some (rdc, rsc, uimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- | _ => returnm (None : option ((mword 3 * mword 3 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
end)
: M (option ((mword 3 * mword 3 * mword 5 * string)))
- else returnm (None : option ((mword 3 * mword 3 * mword 5 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 3 * mword 5 * string))).
-Definition _s2919_ (_s2920_ : string)
-: M (option ((mword 3 * mword 8 * string))) :=
-
- let _s2921_ := _s2920_ in
- (if ((string_startswith _s2921_ "c.addi4spn")) then
- (match (string_drop _s2921_ (projT1 (string_length "c.addi4spn"))) with
- | _s2922_ =>
- (spc_matches_prefix _s2922_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5326_ (_s5327_ : string) : M (option ((mword 3 * mword 8 * string))) :=
+ let _s5328_ := _s5327_ in
+ (if string_startswith _s5328_ "c.addi4spn" then
+ (match (string_drop _s5328_ (projT1 (string_length "c.addi4spn"))) with
+ | _s5329_ =>
+ (spc_matches_prefix _s5329_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2923_ _)) =>
- (match (string_drop _s2922_ _s2923_) with
- | _s2924_ =>
- (creg_name_matches_prefix _s2924_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5330_ _) =>
+ (match (string_drop _s5329_ _s5330_) with
+ | _s5331_ =>
+ (creg_name_matches_prefix _s5331_) >>= fun w__1 : option ((mword 3 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rdc, existT _ _s2925_ _)) =>
- (match (string_drop _s2924_ _s2925_) with
- | _s2926_ =>
- (sep_matches_prefix _s2926_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rdc, existT _ _s5332_ _) =>
+ (match (string_drop _s5331_ _s5332_) with
+ | _s5333_ =>
+ (sep_matches_prefix _s5333_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2927_ _)) =>
- match (string_drop _s2926_ _s2927_) with
- | _s2928_ =>
- match (hex_bits_10_matches_prefix _s2928_) with
- | Some ((v__836, existT _ _s2929_ _)) =>
- if ((eq_vec (subrange_vec_dec v__836 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1))))
- then
- let nzimm : mword 8 := subrange_vec_dec v__836 9 2 in
- let nzimm : mword 8 := subrange_vec_dec v__836 9 2 in
- match (string_drop _s2928_ _s2929_) with
- | s_ => Some ((rdc, nzimm, s_))
- end
- else None
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 3 * mword 8 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s5334_ _) =>
+ match (string_drop _s5333_ _s5334_) with
+ | _s5335_ =>
+ match (hex_bits_10_matches_prefix _s5335_) with
+ | Some (v__1396, existT _ _s5336_ _) =>
+ if eq_vec (subrange_vec_dec v__1396 1 0)
+ ('b"00"
+ : mword (1 - 0 + 1)) then
+ let nzimm : mword 8 := subrange_vec_dec v__1396 9 2 in
+ let nzimm : mword 8 := subrange_vec_dec v__1396 9 2 in
+ match (string_drop _s5335_ _s5336_) with
+ | s_ => Some (rdc, nzimm, s_)
+ end
+ else None
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 3 * mword 8 * string)))
- | _ => returnm (None : option ((mword 3 * mword 8 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8 * string)))
end)
: M (option ((mword 3 * mword 8 * string)))
- | _ => returnm (None : option ((mword 3 * mword 8 * string)))
+ | _ => returnm None
end)
: M (option ((mword 3 * mword 8 * string)))
end)
: M (option ((mword 3 * mword 8 * string)))
- else returnm (None : option ((mword 3 * mword 8 * string))))
+ else returnm None)
: M (option ((mword 3 * mword 8 * string))).
-Definition _s2915_ (_s2916_ : string)
-: option string :=
-
- let _s2917_ := _s2916_ in
- if ((string_startswith _s2917_ "c.nop")) then
- match (string_drop _s2917_ (projT1 (string_length "c.nop"))) with | s_ => Some (s_) end
+Definition _s5322_ (_s5323_ : string) : option string :=
+ let _s5324_ := _s5323_ in
+ if string_startswith _s5324_ "c.nop" then
+ match (string_drop _s5324_ (projT1 (string_length "c.nop"))) with | s_ => Some s_ end
else None.
-Definition _s2891_ (_s2892_ : string)
+Definition _s5296_ (_s5297_ : string)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string))) :=
-
- (match _s2892_ with
- | _s2893_ =>
- (amo_mnemonic_matches_prefix _s2893_) >>= fun w__0 : option ((amoop * {n : Z & ArithFact (n >=
+ (match _s5297_ with
+ | _s5298_ =>
+ (amo_mnemonic_matches_prefix _s5298_) >>= fun w__0 : option ((amoop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2894_ _)) =>
- let _s2895_ := string_drop _s2893_ _s2894_ in
- (if ((string_startswith _s2895_ ".")) then
- (match (string_drop _s2895_ (projT1 (string_length "."))) with
- | _s2896_ =>
- (size_mnemonic_matches_prefix _s2896_) >>= fun w__1 : option ((word_width * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5299_ _) =>
+ let _s5300_ := string_drop _s5298_ _s5299_ in
+ (if string_startswith _s5300_ "." then
+ (match (string_drop _s5300_ (projT1 (string_length "."))) with
+ | _s5301_ =>
+ (size_mnemonic_matches_prefix _s5301_) >>= fun w__1 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((width, existT _ _s2897_ _)) =>
- (match (string_drop _s2896_ _s2897_) with
- | _s2898_ =>
- (maybe_aq_matches_prefix _s2898_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (width, existT _ _s5302_ _) =>
+ (match (string_drop _s5301_ _s5302_) with
+ | _s5303_ =>
+ (maybe_aq_matches_prefix _s5303_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((aq, existT _ _s2899_ _)) =>
- (match (string_drop _s2898_ _s2899_) with
- | _s2900_ =>
- (maybe_rl_matches_prefix _s2900_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s5304_ _) =>
+ (match (string_drop _s5303_ _s5304_) with
+ | _s5305_ =>
+ (maybe_rl_matches_prefix _s5305_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rl, existT _ _s2901_ _)) =>
- (match (string_drop _s2900_ _s2901_) with
- | _s2902_ =>
- (spc_matches_prefix _s2902_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s5306_ _) =>
+ (match (string_drop _s5305_ _s5306_) with
+ | _s5307_ =>
+ (spc_matches_prefix _s5307_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s2903_ _)) =>
- (match (string_drop _s2902_ _s2903_) with
- | _s2904_ =>
- (reg_name_matches_prefix _s2904_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5308_ _) =>
+ (match (string_drop _s5307_ _s5308_) with
+ | _s5309_ =>
+ (reg_name_matches_prefix _s5309_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((rd, existT _ _s2905_ _)) =>
- (match (string_drop _s2904_ _s2905_) with
- | _s2906_ =>
- (sep_matches_prefix _s2906_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5310_ _) =>
+ (match (string_drop _s5309_ _s5310_) with
+ | _s5311_ =>
+ (sep_matches_prefix _s5311_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s2907_ _)) =>
- (match (string_drop _s2906_ _s2907_) with
- | _s2908_ =>
- (reg_name_matches_prefix _s2908_) >>= fun w__7 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5312_ _) =>
+ (match (string_drop _s5311_ _s5312_) with
+ | _s5313_ =>
+ (reg_name_matches_prefix _s5313_) >>= fun w__7 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
- | Some
- ((rs1, existT _ _s2909_ _)) =>
- (match (string_drop _s2908_
- _s2909_) with
- | _s2910_ =>
+ | Some (rs2, existT _ _s5314_ _) =>
+ (match (string_drop _s5313_
+ _s5314_) with
+ | _s5315_ =>
(sep_matches_prefix
- _s2910_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=
+ _s5315_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((tt, existT _ _s2911_ _)) =>
- (match (string_drop
- _s2910_
- _s2911_) with
- | _s2912_ =>
- (reg_name_matches_prefix
- _s2912_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=
- 0)})) =>
- returnm ((match w__9 with
- | Some
- ((rs2, existT _ _s2913_ _)) =>
- match (string_drop
- _s2912_
- _s2913_) with
- | s_ =>
- Some
- ((op, width, aq, rl, rd, rs1, rs2, s_))
- end
- | _ =>
- None
- end)
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- end)
+ (tt, existT _ _s5316_ _) =>
+ let _s5317_ :=
+ string_drop _s5315_
+ _s5316_ in
+ (if string_startswith
+ _s5317_ "(" then
+ (match (string_drop
+ _s5317_
+ (projT1
+ (string_length
+ "("))) with
+ | _s5318_ =>
+ (reg_name_matches_prefix
+ _s5318_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=?
+ 0)})) =>
+ returnm (match w__9 with
+ | Some
+ (rs1, existT _ _s5319_ _) =>
+ let _s5320_ :=
+ string_drop
+ _s5318_
+ _s5319_ in
+ if string_startswith
+ _s5320_
+ ")"
+ then
+ match (string_drop
+ _s5320_
+ (projT1
+ (string_length
+ ")"))) with
+ | s_ =>
+ Some
+ (op, width, aq, rl, rd, rs2, rs1, s_)
+ end
+ else
+ None
+ | _ =>
+ None
+ end)
+ end)
+ : M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ else returnm None)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- else
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string))).
-Definition _s2869_ (_s2870_ : string)
+Definition _s5274_ (_s5275_ : string)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string))) :=
-
- let _s2871_ := _s2870_ in
- (if ((string_startswith _s2871_ "sc.")) then
- (match (string_drop _s2871_ (projT1 (string_length "sc."))) with
- | _s2872_ =>
- (size_mnemonic_matches_prefix _s2872_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s5276_ := _s5275_ in
+ (if string_startswith _s5276_ "sc." then
+ (match (string_drop _s5276_ (projT1 (string_length "sc."))) with
+ | _s5277_ =>
+ (size_mnemonic_matches_prefix _s5277_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s2873_ _)) =>
- (match (string_drop _s2872_ _s2873_) with
- | _s2874_ =>
- (maybe_aq_matches_prefix _s2874_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s5278_ _) =>
+ (match (string_drop _s5277_ _s5278_) with
+ | _s5279_ =>
+ (maybe_aq_matches_prefix _s5279_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s2875_ _)) =>
- (match (string_drop _s2874_ _s2875_) with
- | _s2876_ =>
- (maybe_rl_matches_prefix _s2876_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s5280_ _) =>
+ (match (string_drop _s5279_ _s5280_) with
+ | _s5281_ =>
+ (maybe_rl_matches_prefix _s5281_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s2877_ _)) =>
- (match (string_drop _s2876_ _s2877_) with
- | _s2878_ =>
- (spc_matches_prefix _s2878_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s5282_ _) =>
+ (match (string_drop _s5281_ _s5282_) with
+ | _s5283_ =>
+ (spc_matches_prefix _s5283_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2879_ _)) =>
- (match (string_drop _s2878_ _s2879_) with
- | _s2880_ =>
- (reg_name_matches_prefix _s2880_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5284_ _) =>
+ (match (string_drop _s5283_ _s5284_) with
+ | _s5285_ =>
+ (reg_name_matches_prefix _s5285_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rd, existT _ _s2881_ _)) =>
- (match (string_drop _s2880_ _s2881_) with
- | _s2882_ =>
- (sep_matches_prefix _s2882_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5286_ _) =>
+ (match (string_drop _s5285_ _s5286_) with
+ | _s5287_ =>
+ (sep_matches_prefix _s5287_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2883_ _)) =>
- (match (string_drop _s2882_ _s2883_) with
- | _s2884_ =>
- (reg_name_matches_prefix _s2884_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5288_ _) =>
+ (match (string_drop _s5287_ _s5288_) with
+ | _s5289_ =>
+ (reg_name_matches_prefix _s5289_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((rs1, existT _ _s2885_ _)) =>
- (match (string_drop _s2884_ _s2885_) with
- | _s2886_ =>
- (sep_matches_prefix _s2886_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5290_ _) =>
+ (match (string_drop _s5289_ _s5290_) with
+ | _s5291_ =>
+ (sep_matches_prefix _s5291_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
- | Some ((tt, existT _ _s2887_ _)) =>
- (match (string_drop _s2886_
- _s2887_) with
- | _s2888_ =>
+ | Some (tt, existT _ _s5292_ _) =>
+ (match (string_drop _s5291_
+ _s5292_) with
+ | _s5293_ =>
(reg_name_matches_prefix
- _s2888_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s5293_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__8 with
- | Some
- ((rs2, existT _ _s2889_ _)) =>
- match (string_drop
- _s2888_
- _s2889_) with
- | s_ =>
- Some
- ((size, aq, rl, rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__8 with
+ | Some
+ (rs2, existT _ _s5294_ _) =>
+ match (string_drop
+ _s5293_
+ _s5294_) with
+ | s_ =>
+ Some
+ (size, aq, rl, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)))
- else
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string))).
-Definition _s2851_ (_s2852_ : string)
+Definition _s5256_ (_s5257_ : string)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string))) :=
-
- let _s2853_ := _s2852_ in
- (if ((string_startswith _s2853_ "lr.")) then
- (match (string_drop _s2853_ (projT1 (string_length "lr."))) with
- | _s2854_ =>
- (size_mnemonic_matches_prefix _s2854_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s5258_ := _s5257_ in
+ (if string_startswith _s5258_ "lr." then
+ (match (string_drop _s5258_ (projT1 (string_length "lr."))) with
+ | _s5259_ =>
+ (size_mnemonic_matches_prefix _s5259_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s2855_ _)) =>
- (match (string_drop _s2854_ _s2855_) with
- | _s2856_ =>
- (maybe_aq_matches_prefix _s2856_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s5260_ _) =>
+ (match (string_drop _s5259_ _s5260_) with
+ | _s5261_ =>
+ (maybe_aq_matches_prefix _s5261_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s2857_ _)) =>
- (match (string_drop _s2856_ _s2857_) with
- | _s2858_ =>
- (maybe_rl_matches_prefix _s2858_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s5262_ _) =>
+ (match (string_drop _s5261_ _s5262_) with
+ | _s5263_ =>
+ (maybe_rl_matches_prefix _s5263_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s2859_ _)) =>
- (match (string_drop _s2858_ _s2859_) with
- | _s2860_ =>
- (spc_matches_prefix _s2860_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s5264_ _) =>
+ (match (string_drop _s5263_ _s5264_) with
+ | _s5265_ =>
+ (spc_matches_prefix _s5265_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2861_ _)) =>
- (match (string_drop _s2860_ _s2861_) with
- | _s2862_ =>
- (reg_name_matches_prefix _s2862_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5266_ _) =>
+ (match (string_drop _s5265_ _s5266_) with
+ | _s5267_ =>
+ (reg_name_matches_prefix _s5267_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rd, existT _ _s2863_ _)) =>
- (match (string_drop _s2862_ _s2863_) with
- | _s2864_ =>
- (sep_matches_prefix _s2864_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5268_ _) =>
+ (match (string_drop _s5267_ _s5268_) with
+ | _s5269_ =>
+ (sep_matches_prefix _s5269_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2865_ _)) =>
- (match (string_drop _s2864_ _s2865_) with
- | _s2866_ =>
- (reg_name_matches_prefix _s2866_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5270_ _) =>
+ (match (string_drop _s5269_ _s5270_) with
+ | _s5271_ =>
+ (reg_name_matches_prefix _s5271_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some
- ((rs1, existT _ _s2867_ _)) =>
- match (string_drop _s2866_
- _s2867_) with
- | s_ =>
- Some
- ((size, aq, rl, rd, rs1, s_))
- end
- | _ => None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 5 * string)))
+ returnm (match w__6 with
+ | Some (rs1, existT _ _s5272_ _) =>
+ match (string_drop _s5271_
+ _s5272_) with
+ | s_ =>
+ Some
+ (size, aq, rl, rd, rs1, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((word_width * bool * bool * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string)))
- else returnm (None : option ((word_width * bool * bool * mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 5 * string))).
-Definition _s2839_ (_s2840_ : string)
-: M (option ((mword 5 * mword 5 * string))) :=
-
- let _s2841_ := _s2840_ in
- (if ((string_startswith _s2841_ "sfence.vma")) then
- (match (string_drop _s2841_ (projT1 (string_length "sfence.vma"))) with
- | _s2842_ =>
- (spc_matches_prefix _s2842_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5244_ (_s5245_ : string) : M (option ((mword 5 * mword 5 * string))) :=
+ let _s5246_ := _s5245_ in
+ (if string_startswith _s5246_ "sfence.vma" then
+ (match (string_drop _s5246_ (projT1 (string_length "sfence.vma"))) with
+ | _s5247_ =>
+ (spc_matches_prefix _s5247_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2843_ _)) =>
- (match (string_drop _s2842_ _s2843_) with
- | _s2844_ =>
- (reg_name_matches_prefix _s2844_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5248_ _) =>
+ (match (string_drop _s5247_ _s5248_) with
+ | _s5249_ =>
+ (reg_name_matches_prefix _s5249_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rs1, existT _ _s2845_ _)) =>
- (match (string_drop _s2844_ _s2845_) with
- | _s2846_ =>
- (sep_matches_prefix _s2846_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5250_ _) =>
+ (match (string_drop _s5249_ _s5250_) with
+ | _s5251_ =>
+ (sep_matches_prefix _s5251_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2847_ _)) =>
- (match (string_drop _s2846_ _s2847_) with
- | _s2848_ =>
- (reg_name_matches_prefix _s2848_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5252_ _) =>
+ (match (string_drop _s5251_ _s5252_) with
+ | _s5253_ =>
+ (reg_name_matches_prefix _s5253_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((rs2, existT _ _s2849_ _)) =>
- match (string_drop _s2848_ _s2849_) with
- | s_ => Some ((rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * string)))
+ returnm (match w__3 with
+ | Some (rs2, existT _ _s5254_ _) =>
+ match (string_drop _s5253_ _s5254_) with
+ | s_ => Some (rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * string)))
end)
: M (option ((mword 5 * mword 5 * string)))
- else returnm (None : option ((mword 5 * mword 5 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * string))).
-Definition _s2835_ (_s2836_ : string)
-: option string :=
-
- let _s2837_ := _s2836_ in
- if ((string_startswith _s2837_ "wfi")) then
- match (string_drop _s2837_ (projT1 (string_length "wfi"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s2831_ (_s2832_ : string)
-: option string :=
-
- let _s2833_ := _s2832_ in
- if ((string_startswith _s2833_ "ebreak")) then
- match (string_drop _s2833_ (projT1 (string_length "ebreak"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s2827_ (_s2828_ : string)
-: option string :=
-
- let _s2829_ := _s2828_ in
- if ((string_startswith _s2829_ "sret")) then
- match (string_drop _s2829_ (projT1 (string_length "sret"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s2823_ (_s2824_ : string)
-: option string :=
-
- let _s2825_ := _s2824_ in
- if ((string_startswith _s2825_ "mret")) then
- match (string_drop _s2825_ (projT1 (string_length "mret"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s2819_ (_s2820_ : string)
-: option string :=
-
- let _s2821_ := _s2820_ in
- if ((string_startswith _s2821_ "ecall")) then
- match (string_drop _s2821_ (projT1 (string_length "ecall"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s2815_ (_s2816_ : string)
-: option string :=
-
- let _s2817_ := _s2816_ in
- if ((string_startswith _s2817_ "fence.i")) then
- match (string_drop _s2817_ (projT1 (string_length "fence.i"))) with | s_ => Some (s_) end
- else None.
-
-Definition _s2803_ (_s2804_ : string)
-: M (option ((mword 4 * mword 4 * string))) :=
-
- let _s2805_ := _s2804_ in
- (if ((string_startswith _s2805_ "fence.tso")) then
- (match (string_drop _s2805_ (projT1 (string_length "fence.tso"))) with
- | _s2806_ =>
- (spc_matches_prefix _s2806_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5240_ (_s5241_ : string) : option string :=
+ let _s5242_ := _s5241_ in
+ if string_startswith _s5242_ "wfi" then
+ match (string_drop _s5242_ (projT1 (string_length "wfi"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s5236_ (_s5237_ : string) : option string :=
+ let _s5238_ := _s5237_ in
+ if string_startswith _s5238_ "ebreak" then
+ match (string_drop _s5238_ (projT1 (string_length "ebreak"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s5232_ (_s5233_ : string) : option string :=
+ let _s5234_ := _s5233_ in
+ if string_startswith _s5234_ "sret" then
+ match (string_drop _s5234_ (projT1 (string_length "sret"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s5228_ (_s5229_ : string) : option string :=
+ let _s5230_ := _s5229_ in
+ if string_startswith _s5230_ "mret" then
+ match (string_drop _s5230_ (projT1 (string_length "mret"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s5224_ (_s5225_ : string) : option string :=
+ let _s5226_ := _s5225_ in
+ if string_startswith _s5226_ "ecall" then
+ match (string_drop _s5226_ (projT1 (string_length "ecall"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s5220_ (_s5221_ : string) : option string :=
+ let _s5222_ := _s5221_ in
+ if string_startswith _s5222_ "fence.i" then
+ match (string_drop _s5222_ (projT1 (string_length "fence.i"))) with | s_ => Some s_ end
+ else None.
+
+Definition _s5208_ (_s5209_ : string) : M (option ((mword 4 * mword 4 * string))) :=
+ let _s5210_ := _s5209_ in
+ (if string_startswith _s5210_ "fence.tso" then
+ (match (string_drop _s5210_ (projT1 (string_length "fence.tso"))) with
+ | _s5211_ =>
+ (spc_matches_prefix _s5211_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2807_ _)) =>
- (match (string_drop _s2806_ _s2807_) with
- | _s2808_ =>
- (fence_bits_matches_prefix _s2808_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5212_ _) =>
+ (match (string_drop _s5211_ _s5212_) with
+ | _s5213_ =>
+ (fence_bits_matches_prefix _s5213_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((pred, existT _ _s2809_ _)) =>
- (match (string_drop _s2808_ _s2809_) with
- | _s2810_ =>
- (sep_matches_prefix _s2810_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (pred, existT _ _s5214_ _) =>
+ (match (string_drop _s5213_ _s5214_) with
+ | _s5215_ =>
+ (sep_matches_prefix _s5215_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2811_ _)) =>
- (match (string_drop _s2810_ _s2811_) with
- | _s2812_ =>
- (fence_bits_matches_prefix _s2812_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5216_ _) =>
+ (match (string_drop _s5215_ _s5216_) with
+ | _s5217_ =>
+ (fence_bits_matches_prefix _s5217_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((succ, existT _ _s2813_ _)) =>
- match (string_drop _s2812_ _s2813_) with
- | s_ => Some ((pred, succ, s_))
- end
- | _ => None
- end)
- : option ((mword 4 * mword 4 * string)))
+ returnm (match w__3 with
+ | Some (succ, existT _ _s5218_ _) =>
+ match (string_drop _s5217_ _s5218_) with
+ | s_ => Some (pred, succ, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 4 * mword 4 * string)))
- | _ => returnm (None : option ((mword 4 * mword 4 * string)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4 * string)))
end)
: M (option ((mword 4 * mword 4 * string)))
- | _ => returnm (None : option ((mword 4 * mword 4 * string)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4 * string)))
end)
: M (option ((mword 4 * mword 4 * string)))
- | _ => returnm (None : option ((mword 4 * mword 4 * string)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4 * string)))
end)
: M (option ((mword 4 * mword 4 * string)))
- else returnm (None : option ((mword 4 * mword 4 * string))))
+ else returnm None)
: M (option ((mword 4 * mword 4 * string))).
-Definition _s2791_ (_s2792_ : string)
-: M (option ((mword 4 * mword 4 * string))) :=
-
- let _s2793_ := _s2792_ in
- (if ((string_startswith _s2793_ "fence")) then
- (match (string_drop _s2793_ (projT1 (string_length "fence"))) with
- | _s2794_ =>
- (spc_matches_prefix _s2794_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5196_ (_s5197_ : string) : M (option ((mword 4 * mword 4 * string))) :=
+ let _s5198_ := _s5197_ in
+ (if string_startswith _s5198_ "fence" then
+ (match (string_drop _s5198_ (projT1 (string_length "fence"))) with
+ | _s5199_ =>
+ (spc_matches_prefix _s5199_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2795_ _)) =>
- (match (string_drop _s2794_ _s2795_) with
- | _s2796_ =>
- (fence_bits_matches_prefix _s2796_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5200_ _) =>
+ (match (string_drop _s5199_ _s5200_) with
+ | _s5201_ =>
+ (fence_bits_matches_prefix _s5201_) >>= fun w__1 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((pred, existT _ _s2797_ _)) =>
- (match (string_drop _s2796_ _s2797_) with
- | _s2798_ =>
- (sep_matches_prefix _s2798_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (pred, existT _ _s5202_ _) =>
+ (match (string_drop _s5201_ _s5202_) with
+ | _s5203_ =>
+ (sep_matches_prefix _s5203_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2799_ _)) =>
- (match (string_drop _s2798_ _s2799_) with
- | _s2800_ =>
- (fence_bits_matches_prefix _s2800_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5204_ _) =>
+ (match (string_drop _s5203_ _s5204_) with
+ | _s5205_ =>
+ (fence_bits_matches_prefix _s5205_) >>= fun w__3 : option ((mword 4 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((succ, existT _ _s2801_ _)) =>
- match (string_drop _s2800_ _s2801_) with
- | s_ => Some ((pred, succ, s_))
- end
- | _ => None
- end)
- : option ((mword 4 * mword 4 * string)))
+ returnm (match w__3 with
+ | Some (succ, existT _ _s5206_ _) =>
+ match (string_drop _s5205_ _s5206_) with
+ | s_ => Some (pred, succ, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 4 * mword 4 * string)))
- | _ => returnm (None : option ((mword 4 * mword 4 * string)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4 * string)))
end)
: M (option ((mword 4 * mword 4 * string)))
- | _ => returnm (None : option ((mword 4 * mword 4 * string)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4 * string)))
end)
: M (option ((mword 4 * mword 4 * string)))
- | _ => returnm (None : option ((mword 4 * mword 4 * string)))
+ | _ => returnm None
end)
: M (option ((mword 4 * mword 4 * string)))
end)
: M (option ((mword 4 * mword 4 * string)))
- else returnm (None : option ((mword 4 * mword 4 * string))))
+ else returnm None)
: M (option ((mword 4 * mword 4 * string))).
-Definition _s2774_ (_s2775_ : string)
-: M (option ((sopw * mword 5 * mword 5 * mword 5 * string))) :=
-
- (match _s2775_ with
- | _s2776_ =>
- (shiftiwop_mnemonic_matches_prefix _s2776_) >>= fun w__0 : option ((sopw * {n : Z & ArithFact (n >=
+Definition _s5179_ (_s5180_ : string) : M (option ((sopw * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s5180_ with
+ | _s5181_ =>
+ (shiftiwop_mnemonic_matches_prefix _s5181_) >>= fun w__0 : option ((sopw * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2777_ _)) =>
- (match (string_drop _s2776_ _s2777_) with
- | _s2778_ =>
- (spc_matches_prefix _s2778_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5182_ _) =>
+ (match (string_drop _s5181_ _s5182_) with
+ | _s5183_ =>
+ (spc_matches_prefix _s5183_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2779_ _)) =>
- (match (string_drop _s2778_ _s2779_) with
- | _s2780_ =>
- (reg_name_matches_prefix _s2780_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5184_ _) =>
+ (match (string_drop _s5183_ _s5184_) with
+ | _s5185_ =>
+ (reg_name_matches_prefix _s5185_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2781_ _)) =>
- (match (string_drop _s2780_ _s2781_) with
- | _s2782_ =>
- (sep_matches_prefix _s2782_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5186_ _) =>
+ (match (string_drop _s5185_ _s5186_) with
+ | _s5187_ =>
+ (sep_matches_prefix _s5187_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2783_ _)) =>
- (match (string_drop _s2782_ _s2783_) with
- | _s2784_ =>
- (reg_name_matches_prefix _s2784_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5188_ _) =>
+ (match (string_drop _s5187_ _s5188_) with
+ | _s5189_ =>
+ (reg_name_matches_prefix _s5189_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2785_ _)) =>
- (match (string_drop _s2784_ _s2785_) with
- | _s2786_ =>
- (sep_matches_prefix _s2786_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5190_ _) =>
+ (match (string_drop _s5189_ _s5190_) with
+ | _s5191_ =>
+ (sep_matches_prefix _s5191_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s2787_ _)) =>
- match (string_drop _s2786_ _s2787_) with
- | _s2788_ =>
- match (hex_bits_5_matches_prefix
- _s2788_) with
- | Some ((shamt, existT _ _s2789_ _)) =>
- match (string_drop _s2788_ _s2789_) with
- | s_ =>
- Some ((op, rd, rs1, shamt, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sopw * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s5192_ _) =>
+ match (string_drop _s5191_ _s5192_) with
+ | _s5193_ =>
+ match (hex_bits_5_matches_prefix
+ _s5193_) with
+ | Some (shamt, existT _ _s5194_ _) =>
+ match (string_drop _s5193_ _s5194_) with
+ | s_ =>
+ Some (op, rd, rs1, shamt, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((sopw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((sopw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None : option ((sopw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((sopw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((sopw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sopw * mword 5 * mword 5 * mword 5 * string))).
-Definition _s2757_ (_s2758_ : string)
-: M (option ((ropw * mword 5 * mword 5 * mword 5 * string))) :=
-
- (match _s2758_ with
- | _s2759_ =>
- (rtypew_mnemonic_matches_prefix _s2759_) >>= fun w__0 : option ((ropw * {n : Z & ArithFact (n >=
+Definition _s5162_ (_s5163_ : string) : M (option ((ropw * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s5163_ with
+ | _s5164_ =>
+ (rtypew_mnemonic_matches_prefix _s5164_) >>= fun w__0 : option ((ropw * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2760_ _)) =>
- (match (string_drop _s2759_ _s2760_) with
- | _s2761_ =>
- (spc_matches_prefix _s2761_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5165_ _) =>
+ (match (string_drop _s5164_ _s5165_) with
+ | _s5166_ =>
+ (spc_matches_prefix _s5166_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2762_ _)) =>
- (match (string_drop _s2761_ _s2762_) with
- | _s2763_ =>
- (reg_name_matches_prefix _s2763_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5167_ _) =>
+ (match (string_drop _s5166_ _s5167_) with
+ | _s5168_ =>
+ (reg_name_matches_prefix _s5168_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2764_ _)) =>
- (match (string_drop _s2763_ _s2764_) with
- | _s2765_ =>
- (sep_matches_prefix _s2765_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5169_ _) =>
+ (match (string_drop _s5168_ _s5169_) with
+ | _s5170_ =>
+ (sep_matches_prefix _s5170_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2766_ _)) =>
- (match (string_drop _s2765_ _s2766_) with
- | _s2767_ =>
- (reg_name_matches_prefix _s2767_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5171_ _) =>
+ (match (string_drop _s5170_ _s5171_) with
+ | _s5172_ =>
+ (reg_name_matches_prefix _s5172_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2768_ _)) =>
- (match (string_drop _s2767_ _s2768_) with
- | _s2769_ =>
- (sep_matches_prefix _s2769_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5173_ _) =>
+ (match (string_drop _s5172_ _s5173_) with
+ | _s5174_ =>
+ (sep_matches_prefix _s5174_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2770_ _)) =>
- (match (string_drop _s2769_ _s2770_) with
- | _s2771_ =>
- (reg_name_matches_prefix _s2771_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5175_ _) =>
+ (match (string_drop _s5174_ _s5175_) with
+ | _s5176_ =>
+ (reg_name_matches_prefix _s5176_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s2772_ _)) =>
- match (string_drop _s2771_
- _s2772_) with
- | s_ =>
- Some ((op, rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((ropw * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s5177_ _) =>
+ match (string_drop _s5176_
+ _s5177_) with
+ | s_ =>
+ Some (op, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((ropw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((ropw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((ropw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None : option ((ropw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((ropw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((ropw * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((ropw * mword 5 * mword 5 * mword 5 * string))).
-Definition _s2740_ (_s2741_ : string)
-: M (option ((sop * mword 5 * mword 5 * mword 5 * string))) :=
-
- (match _s2741_ with
- | _s2742_ =>
- (shiftw_mnemonic_matches_prefix _s2742_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=
+Definition _s5145_ (_s5146_ : string) : M (option ((sop * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s5146_ with
+ | _s5147_ =>
+ (shiftw_mnemonic_matches_prefix _s5147_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2743_ _)) =>
- (match (string_drop _s2742_ _s2743_) with
- | _s2744_ =>
- (spc_matches_prefix _s2744_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5148_ _) =>
+ (match (string_drop _s5147_ _s5148_) with
+ | _s5149_ =>
+ (spc_matches_prefix _s5149_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2745_ _)) =>
- (match (string_drop _s2744_ _s2745_) with
- | _s2746_ =>
- (reg_name_matches_prefix _s2746_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5150_ _) =>
+ (match (string_drop _s5149_ _s5150_) with
+ | _s5151_ =>
+ (reg_name_matches_prefix _s5151_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2747_ _)) =>
- (match (string_drop _s2746_ _s2747_) with
- | _s2748_ =>
- (sep_matches_prefix _s2748_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5152_ _) =>
+ (match (string_drop _s5151_ _s5152_) with
+ | _s5153_ =>
+ (sep_matches_prefix _s5153_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2749_ _)) =>
- (match (string_drop _s2748_ _s2749_) with
- | _s2750_ =>
- (reg_name_matches_prefix _s2750_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5154_ _) =>
+ (match (string_drop _s5153_ _s5154_) with
+ | _s5155_ =>
+ (reg_name_matches_prefix _s5155_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2751_ _)) =>
- (match (string_drop _s2750_ _s2751_) with
- | _s2752_ =>
- (sep_matches_prefix _s2752_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5156_ _) =>
+ (match (string_drop _s5155_ _s5156_) with
+ | _s5157_ =>
+ (sep_matches_prefix _s5157_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s2753_ _)) =>
- match (string_drop _s2752_ _s2753_) with
- | _s2754_ =>
- match (hex_bits_5_matches_prefix
- _s2754_) with
- | Some ((shamt, existT _ _s2755_ _)) =>
- match (string_drop _s2754_ _s2755_) with
- | s_ =>
- Some ((op, rd, rs1, shamt, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sop * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s5158_ _) =>
+ match (string_drop _s5157_ _s5158_) with
+ | _s5159_ =>
+ match (hex_bits_5_matches_prefix
+ _s5159_) with
+ | Some (shamt, existT _ _s5160_ _) =>
+ match (string_drop _s5159_ _s5160_) with
+ | s_ =>
+ Some (op, rd, rs1, shamt, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((sop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((sop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None : option ((sop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 5 * string))).
-Definition _s2724_ (_s2725_ : string)
-: M (option ((mword 5 * mword 5 * mword 12 * string))) :=
-
- let _s2726_ := _s2725_ in
- (if ((string_startswith _s2726_ "addiw")) then
- (match (string_drop _s2726_ (projT1 (string_length "addiw"))) with
- | _s2727_ =>
- (spc_matches_prefix _s2727_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s5129_ (_s5130_ : string) : M (option ((mword 5 * mword 5 * mword 12 * string))) :=
+ let _s5131_ := _s5130_ in
+ (if string_startswith _s5131_ "addiw" then
+ (match (string_drop _s5131_ (projT1 (string_length "addiw"))) with
+ | _s5132_ =>
+ (spc_matches_prefix _s5132_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2728_ _)) =>
- (match (string_drop _s2727_ _s2728_) with
- | _s2729_ =>
- (reg_name_matches_prefix _s2729_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5133_ _) =>
+ (match (string_drop _s5132_ _s5133_) with
+ | _s5134_ =>
+ (reg_name_matches_prefix _s5134_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2730_ _)) =>
- (match (string_drop _s2729_ _s2730_) with
- | _s2731_ =>
- (sep_matches_prefix _s2731_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5135_ _) =>
+ (match (string_drop _s5134_ _s5135_) with
+ | _s5136_ =>
+ (sep_matches_prefix _s5136_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2732_ _)) =>
- (match (string_drop _s2731_ _s2732_) with
- | _s2733_ =>
- (reg_name_matches_prefix _s2733_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5137_ _) =>
+ (match (string_drop _s5136_ _s5137_) with
+ | _s5138_ =>
+ (reg_name_matches_prefix _s5138_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s2734_ _)) =>
- (match (string_drop _s2733_ _s2734_) with
- | _s2735_ =>
- (sep_matches_prefix _s2735_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5139_ _) =>
+ (match (string_drop _s5138_ _s5139_) with
+ | _s5140_ =>
+ (sep_matches_prefix _s5140_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2736_ _)) =>
- match (string_drop _s2735_ _s2736_) with
- | _s2737_ =>
- match (hex_bits_12_matches_prefix _s2737_) with
- | Some ((imm, existT _ _s2738_ _)) =>
- match (string_drop _s2737_ _s2738_) with
- | s_ => Some ((rd, rs1, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 12 * string)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s5141_ _) =>
+ match (string_drop _s5140_ _s5141_) with
+ | _s5142_ =>
+ match (hex_bits_12_matches_prefix _s5142_) with
+ | Some (imm, existT _ _s5143_ _) =>
+ match (string_drop _s5142_ _s5143_) with
+ | s_ => Some (rd, rs1, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- else returnm (None : option ((mword 5 * mword 5 * mword 12 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 12 * string))).
-Definition _s2696_ (_s2697_ : string)
+Definition _s5101_ (_s5102_ : string)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string))) :=
-
- let _s2698_ := _s2697_ in
- (if ((string_startswith _s2698_ "s")) then
- (match (string_drop _s2698_ (projT1 (string_length "s"))) with
- | _s2699_ =>
- (size_mnemonic_matches_prefix _s2699_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s5103_ := _s5102_ in
+ (if string_startswith _s5103_ "s" then
+ (match (string_drop _s5103_ (projT1 (string_length "s"))) with
+ | _s5104_ =>
+ (size_mnemonic_matches_prefix _s5104_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s2700_ _)) =>
- (match (string_drop _s2699_ _s2700_) with
- | _s2701_ =>
- (maybe_aq_matches_prefix _s2701_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s5105_ _) =>
+ (match (string_drop _s5104_ _s5105_) with
+ | _s5106_ =>
+ (maybe_aq_matches_prefix _s5106_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((aq, existT _ _s2702_ _)) =>
- (match (string_drop _s2701_ _s2702_) with
- | _s2703_ =>
- (maybe_rl_matches_prefix _s2703_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s5107_ _) =>
+ (match (string_drop _s5106_ _s5107_) with
+ | _s5108_ =>
+ (maybe_rl_matches_prefix _s5108_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rl, existT _ _s2704_ _)) =>
- (match (string_drop _s2703_ _s2704_) with
- | _s2705_ =>
- (spc_matches_prefix _s2705_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s5109_ _) =>
+ (match (string_drop _s5108_ _s5109_) with
+ | _s5110_ =>
+ (spc_matches_prefix _s5110_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2706_ _)) =>
- (match (string_drop _s2705_ _s2706_) with
- | _s2707_ =>
- (reg_name_matches_prefix _s2707_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5111_ _) =>
+ (match (string_drop _s5110_ _s5111_) with
+ | _s5112_ =>
+ (reg_name_matches_prefix _s5112_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs2, existT _ _s2708_ _)) =>
- (match (string_drop _s2707_ _s2708_) with
- | _s2709_ =>
- (sep_matches_prefix _s2709_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s5113_ _) =>
+ (match (string_drop _s5112_ _s5113_) with
+ | _s5114_ =>
+ (sep_matches_prefix _s5114_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2710_ _)) =>
- (match (string_drop _s2709_ _s2710_) with
- | _s2711_ =>
- (match (hex_bits_12_matches_prefix _s2711_) with
- | Some ((imm, existT _ _s2712_ _)) =>
- (match (string_drop _s2711_ _s2712_) with
- | _s2713_ =>
- (opt_spc_matches_prefix _s2713_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5115_ _) =>
+ (match (string_drop _s5114_ _s5115_) with
+ | _s5116_ =>
+ (match (hex_bits_12_matches_prefix _s5116_) with
+ | Some (imm, existT _ _s5117_ _) =>
+ (match (string_drop _s5116_ _s5117_) with
+ | _s5118_ =>
+ (opt_spc_matches_prefix _s5118_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s2714_ _)) =>
- let _s2715_ :=
- string_drop _s2713_ _s2714_ in
- (if ((string_startswith
- _s2715_ "(")) then
- (match (string_drop _s2715_
+ | Some (tt, existT _ _s5119_ _) =>
+ let _s5120_ :=
+ string_drop _s5118_ _s5119_ in
+ (if string_startswith _s5120_
+ "(" then
+ (match (string_drop _s5120_
(projT1
(string_length
"("))) with
- | _s2716_ =>
+ | _s5121_ =>
(opt_spc_matches_prefix
- _s2716_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ _s5121_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
| Some
- ((tt, existT _ _s2717_ _)) =>
+ (tt, existT _ _s5122_ _) =>
(match (string_drop
- _s2716_
- _s2717_) with
- | _s2718_ =>
+ _s5121_
+ _s5122_) with
+ | _s5123_ =>
(reg_name_matches_prefix
- _s2718_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s5123_) >>= fun w__8 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((rs1, existT _ _s2719_ _)) =>
+ (rs1, existT _ _s5124_ _) =>
(match (string_drop
- _s2718_
- _s2719_) with
- | _s2720_ =>
+ _s5123_
+ _s5124_) with
+ | _s5125_ =>
(opt_spc_matches_prefix
- _s2720_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=
+ _s5125_) >>= fun w__9 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__9 with
- | Some
- ((tt, existT _ _s2721_ _)) =>
- let _s2722_ :=
- string_drop
- _s2720_
- _s2721_ in
- if
- ((string_startswith
- _s2722_
- ")"))
- then
- match (string_drop
- _s2722_
- (projT1
- (string_length
- ")"))) with
- | s_ =>
- Some
- ((size, aq, rl, rs2, imm, rs1, s_))
- end
- else
- None
- | _ =>
- None
- end)
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ returnm (match w__9 with
+ | Some
+ (tt, existT _ _s5126_ _) =>
+ let _s5127_ :=
+ string_drop
+ _s5125_
+ _s5126_ in
+ if
+ string_startswith
+ _s5127_
+ ")"
+ then
+ match (string_drop
+ _s5127_
+ (projT1
+ (string_length
+ ")"))) with
+ | s_ =>
+ Some
+ (size, aq, rl, rs2, imm, rs1, s_)
+ end
+ else
+ None
+ | _ =>
+ None
+ end)
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- else
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- else
- returnm (None
- : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string))))
+ else returnm None)
: M (option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string))).
-Definition _s2666_ (_s2667_ : string)
+Definition _s5071_ (_s5072_ : string)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string))) :=
-
- let _s2668_ := _s2667_ in
- (if ((string_startswith _s2668_ "l")) then
- (match (string_drop _s2668_ (projT1 (string_length "l"))) with
- | _s2669_ =>
- (size_mnemonic_matches_prefix _s2669_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=
+ let _s5073_ := _s5072_ in
+ (if string_startswith _s5073_ "l" then
+ (match (string_drop _s5073_ (projT1 (string_length "l"))) with
+ | _s5074_ =>
+ (size_mnemonic_matches_prefix _s5074_) >>= fun w__0 : option ((word_width * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((size, existT _ _s2670_ _)) =>
- (match (string_drop _s2669_ _s2670_) with
- | _s2671_ =>
- (maybe_u_matches_prefix _s2671_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (size, existT _ _s5075_ _) =>
+ (match (string_drop _s5074_ _s5075_) with
+ | _s5076_ =>
+ (maybe_u_matches_prefix _s5076_) >>= fun w__1 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((is_unsigned, existT _ _s2672_ _)) =>
- (match (string_drop _s2671_ _s2672_) with
- | _s2673_ =>
- (maybe_aq_matches_prefix _s2673_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (is_unsigned, existT _ _s5077_ _) =>
+ (match (string_drop _s5076_ _s5077_) with
+ | _s5078_ =>
+ (maybe_aq_matches_prefix _s5078_) >>= fun w__2 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((aq, existT _ _s2674_ _)) =>
- (match (string_drop _s2673_ _s2674_) with
- | _s2675_ =>
- (maybe_rl_matches_prefix _s2675_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=
+ | Some (aq, existT _ _s5079_ _) =>
+ (match (string_drop _s5078_ _s5079_) with
+ | _s5080_ =>
+ (maybe_rl_matches_prefix _s5080_) >>= fun w__3 : option ((bool * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rl, existT _ _s2676_ _)) =>
- (match (string_drop _s2675_ _s2676_) with
- | _s2677_ =>
- (spc_matches_prefix _s2677_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rl, existT _ _s5081_ _) =>
+ (match (string_drop _s5080_ _s5081_) with
+ | _s5082_ =>
+ (spc_matches_prefix _s5082_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((tt, existT _ _s2678_ _)) =>
- (match (string_drop _s2677_ _s2678_) with
- | _s2679_ =>
- (reg_name_matches_prefix _s2679_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5083_ _) =>
+ (match (string_drop _s5082_ _s5083_) with
+ | _s5084_ =>
+ (reg_name_matches_prefix _s5084_) >>= fun w__5 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((rd, existT _ _s2680_ _)) =>
- (match (string_drop _s2679_ _s2680_) with
- | _s2681_ =>
- (sep_matches_prefix _s2681_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5085_ _) =>
+ (match (string_drop _s5084_ _s5085_) with
+ | _s5086_ =>
+ (sep_matches_prefix _s5086_) >>= fun w__6 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__6 with
- | Some ((tt, existT _ _s2682_ _)) =>
- (match (string_drop _s2681_ _s2682_) with
- | _s2683_ =>
+ | Some (tt, existT _ _s5087_ _) =>
+ (match (string_drop _s5086_ _s5087_) with
+ | _s5088_ =>
(match (hex_bits_12_matches_prefix
- _s2683_) with
- | Some
- ((imm, existT _ _s2684_ _)) =>
- (match (string_drop _s2683_
- _s2684_) with
- | _s2685_ =>
+ _s5088_) with
+ | Some (imm, existT _ _s5089_ _) =>
+ (match (string_drop _s5088_
+ _s5089_) with
+ | _s5090_ =>
(opt_spc_matches_prefix
- _s2685_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=
+ _s5090_) >>= fun w__7 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__7 with
| Some
- ((tt, existT _ _s2686_ _)) =>
- let _s2687_ :=
- string_drop _s2685_
- _s2686_ in
- (if ((string_startswith
- _s2687_ "("))
- then
+ (tt, existT _ _s5091_ _) =>
+ let _s5092_ :=
+ string_drop _s5090_
+ _s5091_ in
+ (if string_startswith
+ _s5092_ "(" then
(match (string_drop
- _s2687_
+ _s5092_
(projT1
(string_length
"("))) with
- | _s2688_ =>
+ | _s5093_ =>
(opt_spc_matches_prefix
- _s2688_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=
+ _s5093_) >>= fun w__8 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__8 with
| Some
- ((tt, existT _ _s2689_ _)) =>
+ (tt, existT _ _s5094_ _) =>
(match (string_drop
- _s2688_
- _s2689_) with
- | _s2690_ =>
+ _s5093_
+ _s5094_) with
+ | _s5095_ =>
(reg_name_matches_prefix
- _s2690_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=
+ _s5095_) >>= fun w__9 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__9 with
| Some
- ((rs1, existT _ _s2691_ _)) =>
+ (rs1, existT _ _s5096_ _) =>
(match (string_drop
- _s2690_
- _s2691_) with
- | _s2692_ =>
+ _s5095_
+ _s5096_) with
+ | _s5097_ =>
(opt_spc_matches_prefix
- _s2692_) >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=
+ _s5097_) >>= fun w__10 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__10 with
- | Some
- ((tt, existT _ _s2693_ _)) =>
- let _s2694_ :=
- string_drop
- _s2692_
- _s2693_ in
- if
- ((string_startswith
- _s2694_
- ")"))
- then
- match (string_drop
- _s2694_
- (projT1
- (string_length
- ")"))) with
- | s_ =>
- Some
- ((size, is_unsigned, aq, rl, rd, imm, rs1, s_))
- end
- else
- None
- | _ =>
- None
- end)
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ returnm (match w__10 with
+ | Some
+ (tt, existT _ _s5098_ _) =>
+ let _s5099_ :=
+ string_drop
+ _s5097_
+ _s5098_ in
+ if
+ string_startswith
+ _s5099_
+ ")"
+ then
+ match (string_drop
+ _s5099_
+ (projT1
+ (string_length
+ ")"))) with
+ | s_ =>
+ Some
+ (size, is_unsigned, aq, rl, rd, imm, rs1, s_)
+ end
+ else
+ None
+ | _ =>
+ None
+ end)
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
| _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- else
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string))))
+ else returnm None)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
end)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)))
- else
- returnm (None
- : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string))))
+ else returnm None)
: M (option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string))).
-Definition _s2649_ (_s2650_ : string)
-: M (option ((rop * mword 5 * mword 5 * mword 5 * string))) :=
-
- (match _s2650_ with
- | _s2651_ =>
- (rtype_mnemonic_matches_prefix _s2651_) >>= fun w__0 : option ((rop * {n : Z & ArithFact (n >=
+Definition _s5054_ (_s5055_ : string) : M (option ((rop * mword 5 * mword 5 * mword 5 * string))) :=
+ (match _s5055_ with
+ | _s5056_ =>
+ (rtype_mnemonic_matches_prefix _s5056_) >>= fun w__0 : option ((rop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2652_ _)) =>
- (match (string_drop _s2651_ _s2652_) with
- | _s2653_ =>
- (spc_matches_prefix _s2653_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5057_ _) =>
+ (match (string_drop _s5056_ _s5057_) with
+ | _s5058_ =>
+ (spc_matches_prefix _s5058_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2654_ _)) =>
- (match (string_drop _s2653_ _s2654_) with
- | _s2655_ =>
- (reg_name_matches_prefix _s2655_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5059_ _) =>
+ (match (string_drop _s5058_ _s5059_) with
+ | _s5060_ =>
+ (reg_name_matches_prefix _s5060_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2656_ _)) =>
- (match (string_drop _s2655_ _s2656_) with
- | _s2657_ =>
- (sep_matches_prefix _s2657_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5061_ _) =>
+ (match (string_drop _s5060_ _s5061_) with
+ | _s5062_ =>
+ (sep_matches_prefix _s5062_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2658_ _)) =>
- (match (string_drop _s2657_ _s2658_) with
- | _s2659_ =>
- (reg_name_matches_prefix _s2659_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5063_ _) =>
+ (match (string_drop _s5062_ _s5063_) with
+ | _s5064_ =>
+ (reg_name_matches_prefix _s5064_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2660_ _)) =>
- (match (string_drop _s2659_ _s2660_) with
- | _s2661_ =>
- (sep_matches_prefix _s2661_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5065_ _) =>
+ (match (string_drop _s5064_ _s5065_) with
+ | _s5066_ =>
+ (sep_matches_prefix _s5066_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__5 with
- | Some ((tt, existT _ _s2662_ _)) =>
- (match (string_drop _s2661_ _s2662_) with
- | _s2663_ =>
- (reg_name_matches_prefix _s2663_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5067_ _) =>
+ (match (string_drop _s5066_ _s5067_) with
+ | _s5068_ =>
+ (reg_name_matches_prefix _s5068_) >>= fun w__6 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__6 with
- | Some ((rs2, existT _ _s2664_ _)) =>
- match (string_drop _s2663_
- _s2664_) with
- | s_ =>
- Some ((op, rd, rs1, rs2, s_))
- end
- | _ => None
- end)
- : option ((rop * mword 5 * mword 5 * mword 5 * string)))
+ returnm (match w__6 with
+ | Some (rs2, existT _ _s5069_ _) =>
+ match (string_drop _s5068_
+ _s5069_) with
+ | s_ =>
+ Some (op, rd, rs1, rs2, s_)
+ end
+ | _ => None
+ end)
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((rop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((rop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None
- : option ((rop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
- | _ =>
- returnm (None : option ((rop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((rop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
- | _ => returnm (None : option ((rop * mword 5 * mword 5 * mword 5 * string)))
+ | _ => returnm None
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string)))
end)
: M (option ((rop * mword 5 * mword 5 * mword 5 * string))).
-Definition _s2632_ (_s2633_ : string)
-: M (option ((sop * mword 5 * mword 5 * mword 6 * string))) :=
-
- (match _s2633_ with
- | _s2634_ =>
- (shiftiop_mnemonic_matches_prefix _s2634_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=
+Definition _s5037_ (_s5038_ : string) : M (option ((sop * mword 5 * mword 5 * mword 6 * string))) :=
+ (match _s5038_ with
+ | _s5039_ =>
+ (shiftiop_mnemonic_matches_prefix _s5039_) >>= fun w__0 : option ((sop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2635_ _)) =>
- (match (string_drop _s2634_ _s2635_) with
- | _s2636_ =>
- (spc_matches_prefix _s2636_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5040_ _) =>
+ (match (string_drop _s5039_ _s5040_) with
+ | _s5041_ =>
+ (spc_matches_prefix _s5041_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2637_ _)) =>
- (match (string_drop _s2636_ _s2637_) with
- | _s2638_ =>
- (reg_name_matches_prefix _s2638_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5042_ _) =>
+ (match (string_drop _s5041_ _s5042_) with
+ | _s5043_ =>
+ (reg_name_matches_prefix _s5043_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2639_ _)) =>
- (match (string_drop _s2638_ _s2639_) with
- | _s2640_ =>
- (sep_matches_prefix _s2640_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5044_ _) =>
+ (match (string_drop _s5043_ _s5044_) with
+ | _s5045_ =>
+ (sep_matches_prefix _s5045_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2641_ _)) =>
- (match (string_drop _s2640_ _s2641_) with
- | _s2642_ =>
- (reg_name_matches_prefix _s2642_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5046_ _) =>
+ (match (string_drop _s5045_ _s5046_) with
+ | _s5047_ =>
+ (reg_name_matches_prefix _s5047_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2643_ _)) =>
- (match (string_drop _s2642_ _s2643_) with
- | _s2644_ =>
- (sep_matches_prefix _s2644_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5048_ _) =>
+ (match (string_drop _s5047_ _s5048_) with
+ | _s5049_ =>
+ (sep_matches_prefix _s5049_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s2645_ _)) =>
- match (string_drop _s2644_ _s2645_) with
- | _s2646_ =>
- match (hex_bits_6_matches_prefix
- _s2646_) with
- | Some ((shamt, existT _ _s2647_ _)) =>
- match (string_drop _s2646_ _s2647_) with
- | s_ =>
- Some ((op, rd, rs1, shamt, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((sop * mword 5 * mword 5 * mword 6 * string)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s5050_ _) =>
+ match (string_drop _s5049_ _s5050_) with
+ | _s5051_ =>
+ match (hex_bits_6_matches_prefix
+ _s5051_) with
+ | Some (shamt, existT _ _s5052_ _) =>
+ match (string_drop _s5051_ _s5052_) with
+ | s_ =>
+ Some (op, rd, rs1, shamt, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
- | _ =>
- returnm (None
- : option ((sop * mword 5 * mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
- | _ =>
- returnm (None
- : option ((sop * mword 5 * mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
- | _ =>
- returnm (None : option ((sop * mword 5 * mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
- | _ => returnm (None : option ((sop * mword 5 * mword 5 * mword 6 * string)))
+ | _ => returnm None
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string)))
end)
: M (option ((sop * mword 5 * mword 5 * mword 6 * string))).
-Definition _s2615_ (_s2616_ : string)
-: M (option ((iop * mword 5 * mword 5 * mword 12 * string))) :=
-
- (match _s2616_ with
- | _s2617_ =>
- (itype_mnemonic_matches_prefix _s2617_) >>= fun w__0 : option ((iop * {n : Z & ArithFact (n >=
+Definition _s5020_ (_s5021_ : string) : M (option ((iop * mword 5 * mword 5 * mword 12 * string))) :=
+ (match _s5021_ with
+ | _s5022_ =>
+ (itype_mnemonic_matches_prefix _s5022_) >>= fun w__0 : option ((iop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2618_ _)) =>
- (match (string_drop _s2617_ _s2618_) with
- | _s2619_ =>
- (spc_matches_prefix _s2619_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5023_ _) =>
+ (match (string_drop _s5022_ _s5023_) with
+ | _s5024_ =>
+ (spc_matches_prefix _s5024_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2620_ _)) =>
- (match (string_drop _s2619_ _s2620_) with
- | _s2621_ =>
- (reg_name_matches_prefix _s2621_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5025_ _) =>
+ (match (string_drop _s5024_ _s5025_) with
+ | _s5026_ =>
+ (reg_name_matches_prefix _s5026_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2622_ _)) =>
- (match (string_drop _s2621_ _s2622_) with
- | _s2623_ =>
- (sep_matches_prefix _s2623_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s5027_ _) =>
+ (match (string_drop _s5026_ _s5027_) with
+ | _s5028_ =>
+ (sep_matches_prefix _s5028_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2624_ _)) =>
- (match (string_drop _s2623_ _s2624_) with
- | _s2625_ =>
- (reg_name_matches_prefix _s2625_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5029_ _) =>
+ (match (string_drop _s5028_ _s5029_) with
+ | _s5030_ =>
+ (reg_name_matches_prefix _s5030_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs1, existT _ _s2626_ _)) =>
- (match (string_drop _s2625_ _s2626_) with
- | _s2627_ =>
- (sep_matches_prefix _s2627_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5031_ _) =>
+ (match (string_drop _s5030_ _s5031_) with
+ | _s5032_ =>
+ (sep_matches_prefix _s5032_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s2628_ _)) =>
- match (string_drop _s2627_ _s2628_) with
- | _s2629_ =>
- match (hex_bits_12_matches_prefix
- _s2629_) with
- | Some ((imm, existT _ _s2630_ _)) =>
- match (string_drop _s2629_ _s2630_) with
- | s_ =>
- Some ((op, rd, rs1, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((iop * mword 5 * mword 5 * mword 12 * string)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s5033_ _) =>
+ match (string_drop _s5032_ _s5033_) with
+ | _s5034_ =>
+ match (hex_bits_12_matches_prefix
+ _s5034_) with
+ | Some (imm, existT _ _s5035_ _) =>
+ match (string_drop _s5034_ _s5035_) with
+ | s_ => Some (op, rd, rs1, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((iop * mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((iop * mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None : option ((iop * mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((iop * mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((iop * mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((iop * mword 5 * mword 5 * mword 12 * string))).
-Definition _s2598_ (_s2599_ : string)
-: M (option ((bop * mword 5 * mword 5 * mword 13 * string))) :=
-
- (match _s2599_ with
- | _s2600_ =>
- (btype_mnemonic_matches_prefix _s2600_) >>= fun w__0 : option ((bop * {n : Z & ArithFact (n >=
+Definition _s5003_ (_s5004_ : string) : M (option ((bop * mword 5 * mword 5 * mword 13 * string))) :=
+ (match _s5004_ with
+ | _s5005_ =>
+ (btype_mnemonic_matches_prefix _s5005_) >>= fun w__0 : option ((bop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2601_ _)) =>
- (match (string_drop _s2600_ _s2601_) with
- | _s2602_ =>
- (spc_matches_prefix _s2602_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s5006_ _) =>
+ (match (string_drop _s5005_ _s5006_) with
+ | _s5007_ =>
+ (spc_matches_prefix _s5007_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2603_ _)) =>
- (match (string_drop _s2602_ _s2603_) with
- | _s2604_ =>
- (reg_name_matches_prefix _s2604_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5008_ _) =>
+ (match (string_drop _s5007_ _s5008_) with
+ | _s5009_ =>
+ (reg_name_matches_prefix _s5009_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rs1, existT _ _s2605_ _)) =>
- (match (string_drop _s2604_ _s2605_) with
- | _s2606_ =>
- (sep_matches_prefix _s2606_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s5010_ _) =>
+ (match (string_drop _s5009_ _s5010_) with
+ | _s5011_ =>
+ (sep_matches_prefix _s5011_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((tt, existT _ _s2607_ _)) =>
- (match (string_drop _s2606_ _s2607_) with
- | _s2608_ =>
- (reg_name_matches_prefix _s2608_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s5012_ _) =>
+ (match (string_drop _s5011_ _s5012_) with
+ | _s5013_ =>
+ (reg_name_matches_prefix _s5013_) >>= fun w__4 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__4 with
- | Some ((rs2, existT _ _s2609_ _)) =>
- (match (string_drop _s2608_ _s2609_) with
- | _s2610_ =>
- (sep_matches_prefix _s2610_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs2, existT _ _s5014_ _) =>
+ (match (string_drop _s5013_ _s5014_) with
+ | _s5015_ =>
+ (sep_matches_prefix _s5015_) >>= fun w__5 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__5 with
- | Some ((tt, existT _ _s2611_ _)) =>
- match (string_drop _s2610_ _s2611_) with
- | _s2612_ =>
- match (hex_bits_13_matches_prefix
- _s2612_) with
- | Some ((imm, existT _ _s2613_ _)) =>
- match (string_drop _s2612_ _s2613_) with
- | s_ =>
- Some ((op, rs1, rs2, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((bop * mword 5 * mword 5 * mword 13 * string)))
+ returnm (match w__5 with
+ | Some (tt, existT _ _s5016_ _) =>
+ match (string_drop _s5015_ _s5016_) with
+ | _s5017_ =>
+ match (hex_bits_13_matches_prefix
+ _s5017_) with
+ | Some (imm, existT _ _s5018_ _) =>
+ match (string_drop _s5017_ _s5018_) with
+ | s_ =>
+ Some (op, rs1, rs2, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
- | _ =>
- returnm (None
- : option ((bop * mword 5 * mword 5 * mword 13 * string)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
- | _ =>
- returnm (None
- : option ((bop * mword 5 * mword 5 * mword 13 * string)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
- | _ =>
- returnm (None : option ((bop * mword 5 * mword 5 * mword 13 * string)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
- | _ => returnm (None : option ((bop * mword 5 * mword 5 * mword 13 * string)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
- | _ => returnm (None : option ((bop * mword 5 * mword 5 * mword 13 * string)))
+ | _ => returnm None
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string)))
end)
: M (option ((bop * mword 5 * mword 5 * mword 13 * string))).
-Definition _s2582_ (_s2583_ : string)
-: M (option ((mword 5 * mword 5 * mword 12 * string))) :=
-
- let _s2584_ := _s2583_ in
- (if ((string_startswith _s2584_ "jalr")) then
- (match (string_drop _s2584_ (projT1 (string_length "jalr"))) with
- | _s2585_ =>
- (spc_matches_prefix _s2585_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s4987_ (_s4988_ : string) : M (option ((mword 5 * mword 5 * mword 12 * string))) :=
+ let _s4989_ := _s4988_ in
+ (if string_startswith _s4989_ "jalr" then
+ (match (string_drop _s4989_ (projT1 (string_length "jalr"))) with
+ | _s4990_ =>
+ (spc_matches_prefix _s4990_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2586_ _)) =>
- (match (string_drop _s2585_ _s2586_) with
- | _s2587_ =>
- (reg_name_matches_prefix _s2587_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s4991_ _) =>
+ (match (string_drop _s4990_ _s4991_) with
+ | _s4992_ =>
+ (reg_name_matches_prefix _s4992_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2588_ _)) =>
- (match (string_drop _s2587_ _s2588_) with
- | _s2589_ =>
- (sep_matches_prefix _s2589_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s4993_ _) =>
+ (match (string_drop _s4992_ _s4993_) with
+ | _s4994_ =>
+ (sep_matches_prefix _s4994_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((tt, existT _ _s2590_ _)) =>
- (match (string_drop _s2589_ _s2590_) with
- | _s2591_ =>
- (reg_name_matches_prefix _s2591_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s4995_ _) =>
+ (match (string_drop _s4994_ _s4995_) with
+ | _s4996_ =>
+ (reg_name_matches_prefix _s4996_) >>= fun w__3 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__3 with
- | Some ((rs1, existT _ _s2592_ _)) =>
- (match (string_drop _s2591_ _s2592_) with
- | _s2593_ =>
- (sep_matches_prefix _s2593_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rs1, existT _ _s4997_ _) =>
+ (match (string_drop _s4996_ _s4997_) with
+ | _s4998_ =>
+ (sep_matches_prefix _s4998_) >>= fun w__4 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__4 with
- | Some ((tt, existT _ _s2594_ _)) =>
- match (string_drop _s2593_ _s2594_) with
- | _s2595_ =>
- match (hex_bits_12_matches_prefix _s2595_) with
- | Some ((imm, existT _ _s2596_ _)) =>
- match (string_drop _s2595_ _s2596_) with
- | s_ => Some ((rd, rs1, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 5 * mword 12 * string)))
+ returnm (match w__4 with
+ | Some (tt, existT _ _s4999_ _) =>
+ match (string_drop _s4998_ _s4999_) with
+ | _s5000_ =>
+ match (hex_bits_12_matches_prefix _s5000_) with
+ | Some (imm, existT _ _s5001_ _) =>
+ match (string_drop _s5000_ _s5001_) with
+ | s_ => Some (rd, rs1, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- | _ =>
- returnm (None
- : option ((mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- | _ => returnm (None : option ((mword 5 * mword 5 * mword 12 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
end)
: M (option ((mword 5 * mword 5 * mword 12 * string)))
- else returnm (None : option ((mword 5 * mword 5 * mword 12 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 5 * mword 12 * string))).
-Definition _s2570_ (_s2571_ : string)
-: M (option ((mword 5 * mword 21 * string))) :=
-
- let _s2572_ := _s2571_ in
- (if ((string_startswith _s2572_ "jal")) then
- (match (string_drop _s2572_ (projT1 (string_length "jal"))) with
- | _s2573_ =>
- (spc_matches_prefix _s2573_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >= 0)})) =>
+Definition _s4975_ (_s4976_ : string) : M (option ((mword 5 * mword 21 * string))) :=
+ let _s4977_ := _s4976_ in
+ (if string_startswith _s4977_ "jal" then
+ (match (string_drop _s4977_ (projT1 (string_length "jal"))) with
+ | _s4978_ =>
+ (spc_matches_prefix _s4978_) >>= fun w__0 : option ((unit * {n : Z & ArithFact (n >=? 0)})) =>
(match w__0 with
- | Some ((tt, existT _ _s2574_ _)) =>
- (match (string_drop _s2573_ _s2574_) with
- | _s2575_ =>
- (reg_name_matches_prefix _s2575_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s4979_ _) =>
+ (match (string_drop _s4978_ _s4979_) with
+ | _s4980_ =>
+ (reg_name_matches_prefix _s4980_) >>= fun w__1 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((rd, existT _ _s2576_ _)) =>
- (match (string_drop _s2575_ _s2576_) with
- | _s2577_ =>
- (sep_matches_prefix _s2577_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s4981_ _) =>
+ (match (string_drop _s4980_ _s4981_) with
+ | _s4982_ =>
+ (sep_matches_prefix _s4982_) >>= fun w__2 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__2 with
- | Some ((tt, existT _ _s2578_ _)) =>
- match (string_drop _s2577_ _s2578_) with
- | _s2579_ =>
- match (hex_bits_21_matches_prefix _s2579_) with
- | Some ((imm, existT _ _s2580_ _)) =>
- match (string_drop _s2579_ _s2580_) with
- | s_ => Some ((rd, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((mword 5 * mword 21 * string)))
+ returnm (match w__2 with
+ | Some (tt, existT _ _s4983_ _) =>
+ match (string_drop _s4982_ _s4983_) with
+ | _s4984_ =>
+ match (hex_bits_21_matches_prefix _s4984_) with
+ | Some (imm, existT _ _s4985_ _) =>
+ match (string_drop _s4984_ _s4985_) with
+ | s_ => Some (rd, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((mword 5 * mword 21 * string)))
- | _ => returnm (None : option ((mword 5 * mword 21 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 21 * string)))
end)
: M (option ((mword 5 * mword 21 * string)))
- | _ => returnm (None : option ((mword 5 * mword 21 * string)))
+ | _ => returnm None
end)
: M (option ((mword 5 * mword 21 * string)))
end)
: M (option ((mword 5 * mword 21 * string)))
- else returnm (None : option ((mword 5 * mword 21 * string))))
+ else returnm None)
: M (option ((mword 5 * mword 21 * string))).
-Definition _s2557_ (_s2558_ : string)
-: M (option ((uop * mword 5 * mword 20 * string))) :=
-
- (match _s2558_ with
- | _s2559_ =>
- (utype_mnemonic_matches_prefix _s2559_) >>= fun w__0 : option ((uop * {n : Z & ArithFact (n >=
+Definition _s4962_ (_s4963_ : string) : M (option ((uop * mword 5 * mword 20 * string))) :=
+ (match _s4963_ with
+ | _s4964_ =>
+ (utype_mnemonic_matches_prefix _s4964_) >>= fun w__0 : option ((uop * {n : Z & ArithFact (n >=?
0)})) =>
(match w__0 with
- | Some ((op, existT _ _s2560_ _)) =>
- (match (string_drop _s2559_ _s2560_) with
- | _s2561_ =>
- (spc_matches_prefix _s2561_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (op, existT _ _s4965_ _) =>
+ (match (string_drop _s4964_ _s4965_) with
+ | _s4966_ =>
+ (spc_matches_prefix _s4966_) >>= fun w__1 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
(match w__1 with
- | Some ((tt, existT _ _s2562_ _)) =>
- (match (string_drop _s2561_ _s2562_) with
- | _s2563_ =>
- (reg_name_matches_prefix _s2563_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=
+ | Some (tt, existT _ _s4967_ _) =>
+ (match (string_drop _s4966_ _s4967_) with
+ | _s4968_ =>
+ (reg_name_matches_prefix _s4968_) >>= fun w__2 : option ((mword 5 * {n : Z & ArithFact (n >=?
0)})) =>
(match w__2 with
- | Some ((rd, existT _ _s2564_ _)) =>
- (match (string_drop _s2563_ _s2564_) with
- | _s2565_ =>
- (sep_matches_prefix _s2565_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=
+ | Some (rd, existT _ _s4969_ _) =>
+ (match (string_drop _s4968_ _s4969_) with
+ | _s4970_ =>
+ (sep_matches_prefix _s4970_) >>= fun w__3 : option ((unit * {n : Z & ArithFact (n >=?
0)})) =>
- returnm ((match w__3 with
- | Some ((tt, existT _ _s2566_ _)) =>
- match (string_drop _s2565_ _s2566_) with
- | _s2567_ =>
- match (hex_bits_20_matches_prefix _s2567_) with
- | Some ((imm, existT _ _s2568_ _)) =>
- match (string_drop _s2567_ _s2568_) with
- | s_ => Some ((op, rd, imm, s_))
- end
- | _ => None
- end
- end
- | _ => None
- end)
- : option ((uop * mword 5 * mword 20 * string)))
+ returnm (match w__3 with
+ | Some (tt, existT _ _s4971_ _) =>
+ match (string_drop _s4970_ _s4971_) with
+ | _s4972_ =>
+ match (hex_bits_20_matches_prefix _s4972_) with
+ | Some (imm, existT _ _s4973_ _) =>
+ match (string_drop _s4972_ _s4973_) with
+ | s_ => Some (op, rd, imm, s_)
+ end
+ | _ => None
+ end
+ end
+ | _ => None
+ end)
end)
: M (option ((uop * mword 5 * mword 20 * string)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20 * string)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20 * string)))
end)
: M (option ((uop * mword 5 * mword 20 * string)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20 * string)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20 * string)))
end)
: M (option ((uop * mword 5 * mword 20 * string)))
- | _ => returnm (None : option ((uop * mword 5 * mword 20 * string)))
+ | _ => returnm None
end)
: M (option ((uop * mword 5 * mword 20 * string)))
end)
: M (option ((uop * mword 5 * mword 20 * string))).
-Definition assembly_matches_prefix (arg_ : string)
-: M (option ((ast * {n : Z & ArithFact (n >= 0)}))) :=
-
- let _s2569_ := arg_ in
- (_s2557_ _s2569_) >>= fun w__0 : option ((uop * mword 5 * mword 20 * string)) =>
- (if ((match w__0 with | Some ((op, rd, imm, s_)) => true | _ => false end)) then
- (_s2557_ _s2569_) >>= fun w__1 : option ((uop * mword 5 * mword 20 * string)) =>
+Definition assembly_matches_prefix (arg_ : string)
+: M (option ((ast * {n : Z & ArithFact (n >=? 0)}))) :=
+ let _s4974_ := arg_ in
+ (_s4962_ _s4974_) >>= fun w__0 : option ((uop * mword 5 * mword 20 * string)) =>
+ (if match w__0 with | Some (op, rd, imm, s_) => true | _ => false end then
+ (_s4962_ _s4974_) >>= fun w__1 : option ((uop * mword 5 * mword 20 * string)) =>
(match w__1 with
- | Some ((op, rd, imm, s_)) =>
- returnm ((Some
- ((UTYPE
- ((imm, rd, op)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (op, rd, imm, s_) =>
+ returnm (Some
+ (UTYPE (imm, rd, op), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2570_ _s2569_) >>= fun w__4 : option ((mword 5 * mword 21 * string)) =>
- (if ((match w__4 with | Some ((rd, imm, s_)) => true | _ => false end)) then
- (_s2570_ _s2569_) >>= fun w__5 : option ((mword 5 * mword 21 * string)) =>
+ (_s4975_ _s4974_) >>= fun w__4 : option ((mword 5 * mword 21 * string)) =>
+ (if match w__4 with | Some (rd, imm, s_) => true | _ => false end then
+ (_s4975_ _s4974_) >>= fun w__5 : option ((mword 5 * mword 21 * string)) =>
(match w__5 with
- | Some ((rd, imm, s_)) =>
- returnm ((Some
- ((RISCV_JAL
- ((imm, rd)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (rd, imm, s_) =>
+ returnm (Some
+ (RISCV_JAL (imm, rd), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2582_ _s2569_) >>= fun w__8 : option ((mword 5 * mword 5 * mword 12 * string)) =>
- (if ((match w__8 with | Some ((rd, rs1, imm, s_)) => true | _ => false end)) then
- (_s2582_ _s2569_) >>= fun w__9 : option ((mword 5 * mword 5 * mword 12 * string)) =>
+ (_s4987_ _s4974_) >>= fun w__8 : option ((mword 5 * mword 5 * mword 12 * string)) =>
+ (if match w__8 with | Some (rd, rs1, imm, s_) => true | _ => false end then
+ (_s4987_ _s4974_) >>= fun w__9 : option ((mword 5 * mword 5 * mword 12 * string)) =>
(match w__9 with
- | Some ((rd, rs1, imm, s_)) =>
- returnm ((Some
- ((RISCV_JALR
- ((imm, rs1, rd)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (rd, rs1, imm, s_) =>
+ returnm (Some
+ (RISCV_JALR (imm, rs1, rd), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2598_ _s2569_) >>= fun w__12 : option ((bop * mword 5 * mword 5 * mword 13 * string)) =>
- (if ((match w__12 with | Some ((op, rs1, rs2, imm, s_)) => true | _ => false end)) then
- (_s2598_ _s2569_) >>= fun w__13 : option ((bop * mword 5 * mword 5 * mword 13 * string)) =>
+ (_s5003_ _s4974_) >>= fun w__12 : option ((bop * mword 5 * mword 5 * mword 13 * string)) =>
+ (if match w__12 with | Some (op, rs1, rs2, imm, s_) => true | _ => false end then
+ (_s5003_ _s4974_) >>= fun w__13 : option ((bop * mword 5 * mword 5 * mword 13 * string)) =>
(match w__13 with
- | Some ((op, rs1, rs2, imm, s_)) =>
- returnm ((Some
- ((BTYPE
- ((imm, rs2, rs1, op)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (op, rs1, rs2, imm, s_) =>
+ returnm (Some
+ (BTYPE (imm, rs2, rs1, op), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2615_ _s2569_) >>= fun w__16 : option ((iop * mword 5 * mword 5 * mword 12 * string)) =>
- (if ((match w__16 with | Some ((op, rd, rs1, imm, s_)) => true | _ => false end))
- then
- (_s2615_ _s2569_) >>= fun w__17 : option ((iop * mword 5 * mword 5 * mword 12 * string)) =>
+ (_s5020_ _s4974_) >>= fun w__16 : option ((iop * mword 5 * mword 5 * mword 12 * string)) =>
+ (if match w__16 with | Some (op, rd, rs1, imm, s_) => true | _ => false end then
+ (_s5020_ _s4974_) >>= fun w__17 : option ((iop * mword 5 * mword 5 * mword 12 * string)) =>
(match w__17 with
- | Some ((op, rd, rs1, imm, s_)) =>
- returnm ((Some
- ((ITYPE
- ((imm, rs1, rd, op)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (op, rd, rs1, imm, s_) =>
+ returnm (Some
+ (ITYPE (imm, rs1, rd, op), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_)) (projT1 (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2632_ _s2569_) >>= fun w__20 : option ((sop * mword 5 * mword 5 * mword 6 * string)) =>
- (if ((match w__20 with | Some ((op, rd, rs1, shamt, s_)) => true | _ => false end))
- then
- (_s2632_ _s2569_) >>= fun w__21 : option ((sop * mword 5 * mword 5 * mword 6 * string)) =>
+ (_s5037_ _s4974_) >>= fun w__20 : option ((sop * mword 5 * mword 5 * mword 6 * string)) =>
+ (if match w__20 with | Some (op, rd, rs1, shamt, s_) => true | _ => false end then
+ (_s5037_ _s4974_) >>= fun w__21 : option ((sop * mword 5 * mword 5 * mword 6 * string)) =>
(match w__21 with
- | Some ((op, rd, rs1, shamt, s_)) =>
- returnm ((Some
- ((SHIFTIOP
- ((shamt, rs1, rd, op)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (op, rd, rs1, shamt, s_) =>
+ returnm (Some
+ (SHIFTIOP (shamt, rs1, rd, op), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2649_ _s2569_) >>= fun w__24 : option ((rop * mword 5 * mword 5 * mword 5 * string)) =>
- (if ((match w__24 with | Some ((op, rd, rs1, rs2, s_)) => true | _ => false end))
+ (_s5054_ _s4974_) >>= fun w__24 : option ((rop * mword 5 * mword 5 * mword 5 * string)) =>
+ (if match w__24 with | Some (op, rd, rs1, rs2, s_) => true | _ => false end
then
- (_s2649_ _s2569_) >>= fun w__25 : option ((rop * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5054_ _s4974_) >>= fun w__25 : option ((rop * mword 5 * mword 5 * mword 5 * string)) =>
(match w__25 with
- | Some ((op, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((RTYPE
- ((rs2, rs1, rd, op)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (op, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (RTYPE (rs2, rs1, rd, op), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2666_ _s2569_) >>= fun w__28 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)) =>
- (if ((match w__28 with
- | Some ((size, is_unsigned, aq, rl, rd, imm, rs1, s_)) => true
- | _ => false
- end)) then
- (_s2666_ _s2569_) >>= fun w__29 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)) =>
+ (_s5071_ _s4974_) >>= fun w__28 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)) =>
+ (if match w__28 with
+ | Some (size, is_unsigned, aq, rl, rd, imm, rs1, s_) => true
+ | _ => false
+ end then
+ (_s5071_ _s4974_) >>= fun w__29 : option ((word_width * bool * bool * bool * mword 5 * mword 12 * mword 5 * string)) =>
(match w__29 with
- | Some ((size, is_unsigned, aq, rl, rd, imm, rs1, s_)) =>
- returnm ((Some
- ((LOAD
- ((imm, rs1, rd, is_unsigned, size, aq, rl)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (size, is_unsigned, aq, rl, rd, imm, rs1, s_) =>
+ returnm (Some
+ (LOAD (imm, rs1, rd, is_unsigned, size, aq, rl), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2696_ _s2569_) >>= fun w__32 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)) =>
- (if ((match w__32 with
- | Some ((size, aq, rl, rs2, imm, rs1, s_)) => true
- | _ => false
- end)) then
- (_s2696_ _s2569_) >>= fun w__33 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)) =>
+ (_s5101_ _s4974_) >>= fun w__32 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)) =>
+ (if match w__32 with
+ | Some (size, aq, rl, rs2, imm, rs1, s_) => true
+ | _ => false
+ end then
+ (_s5101_ _s4974_) >>= fun w__33 : option ((word_width * bool * bool * mword 5 * mword 12 * mword 5 * string)) =>
(match w__33 with
- | Some ((size, aq, rl, rs2, imm, rs1, s_)) =>
- returnm ((Some
- ((STORE
- ((imm, rs2, rs1, size, aq, rl)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
- | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (size, aq, rl, rs2, imm, rs1, s_) =>
+ returnm (Some
+ (STORE (imm, rs2, rs1, size, aq, rl), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
+ | _ => exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2724_ _s2569_) >>= fun w__36 : option ((mword 5 * mword 5 * mword 12 * string)) =>
- (if ((match w__36 with
- | Some ((rd, rs1, imm, s_)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s2724_ _s2569_) >>= fun w__37 : option ((mword 5 * mword 5 * mword 12 * string)) =>
+ (_s5129_ _s4974_) >>= fun w__36 : option ((mword 5 * mword 5 * mword 12 * string)) =>
+ (if match w__36 with
+ | Some (rd, rs1, imm, s_) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s5129_ _s4974_) >>= fun w__37 : option ((mword 5 * mword 5 * mword 12 * string)) =>
(match w__37 with
- | Some ((rd, rs1, imm, s_)) =>
- returnm ((Some
- ((ADDIW
- ((imm, rs1, rd)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (rd, rs1, imm, s_) =>
+ returnm (Some
+ (ADDIW (imm, rs1, rd), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
- exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ exit tt : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2740_ _s2569_) >>= fun w__40 : option ((sop * mword 5 * mword 5 * mword 5 * string)) =>
- (if ((match w__40 with
- | Some ((op, rd, rs1, shamt, s_)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s2740_ _s2569_) >>= fun w__41 : option ((sop * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5145_ _s4974_) >>= fun w__40 : option ((sop * mword 5 * mword 5 * mword 5 * string)) =>
+ (if match w__40 with
+ | Some (op, rd, rs1, shamt, s_) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s5145_ _s4974_) >>= fun w__41 : option ((sop * mword 5 * mword 5 * mword 5 * string)) =>
(match w__41 with
- | Some ((op, rd, rs1, shamt, s_)) =>
- returnm ((Some
- ((SHIFTW
- ((shamt, rs1, rd, op)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (op, rd, rs1, shamt, s_) =>
+ returnm (Some
+ (SHIFTW (shamt, rs1, rd, op), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
- exit tt : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2757_ _s2569_) >>= fun w__44 : option ((ropw * mword 5 * mword 5 * mword 5 * string)) =>
- (if ((match w__44 with
- | Some ((op, rd, rs1, rs2, s_)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s2757_ _s2569_) >>= fun w__45 : option ((ropw * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5162_ _s4974_) >>= fun w__44 : option ((ropw * mword 5 * mword 5 * mword 5 * string)) =>
+ (if match w__44 with
+ | Some (op, rd, rs1, rs2, s_) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s5162_ _s4974_) >>= fun w__45 : option ((ropw * mword 5 * mword 5 * mword 5 * string)) =>
(match w__45 with
- | Some ((op, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((RTYPEW
- ((rs2, rs1, rd, op)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (op, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (RTYPEW (rs2, rs1, rd, op), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2774_ _s2569_) >>= fun w__48 : option ((sopw * mword 5 * mword 5 * mword 5 * string)) =>
- (if ((match w__48 with
- | Some ((op, rd, rs1, shamt, s_)) => Z.eqb 64 64
- | _ => false
- end)) then
- (_s2774_ _s2569_) >>= fun w__49 : option ((sopw * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5179_ _s4974_) >>= fun w__48 : option ((sopw * mword 5 * mword 5 * mword 5 * string)) =>
+ (if match w__48 with
+ | Some (op, rd, rs1, shamt, s_) => Z.eqb 64 64
+ | _ => false
+ end then
+ (_s5179_ _s4974_) >>= fun w__49 : option ((sopw * mword 5 * mword 5 * mword 5 * string)) =>
(match w__49 with
- | Some ((op, rd, rs1, shamt, s_)) =>
- returnm ((Some
- ((SHIFTIWOP
- ((shamt, rs1, rd, op)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (op, rd, rs1, shamt, s_) =>
+ returnm (Some
+ (SHIFTIWOP (shamt, rs1, rd, op), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2791_ _s2569_) >>= fun w__52 : option ((mword 4 * mword 4 * string)) =>
- (if ((match w__52 with
- | Some ((pred, succ, s_)) => true
- | _ => false
- end)) then
- (_s2791_ _s2569_) >>= fun w__53 : option ((mword 4 * mword 4 * string)) =>
+ (_s5196_ _s4974_) >>= fun w__52 : option ((mword 4 * mword 4 * string)) =>
+ (if match w__52 with
+ | Some (pred, succ, s_) => true
+ | _ => false
+ end then
+ (_s5196_ _s4974_) >>= fun w__53 : option ((mword 4 * mword 4 * string)) =>
(match w__53 with
- | Some ((pred, succ, s_)) =>
- returnm ((Some
- ((FENCE
- ((pred, succ)), build_ex
- (projT1
- (sub_nat (projT1 (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (pred, succ, s_) =>
+ returnm (Some
+ (FENCE (pred, succ), build_ex
+ (projT1
+ (sub_nat (projT1 (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2803_ _s2569_) >>= fun w__56 : option ((mword 4 * mword 4 * string)) =>
- (if ((match w__56 with
- | Some ((pred, succ, s_)) => true
- | _ => false
- end)) then
- (_s2803_ _s2569_) >>= fun w__57 : option ((mword 4 * mword 4 * string)) =>
+ (_s5208_ _s4974_) >>= fun w__56 : option ((mword 4 * mword 4 * string)) =>
+ (if match w__56 with
+ | Some (pred, succ, s_) => true
+ | _ => false
+ end then
+ (_s5208_ _s4974_) >>= fun w__57 : option ((mword 4 * mword 4 * string)) =>
(match w__57 with
- | Some ((pred, succ, s_)) =>
- returnm ((Some
- ((FENCE_TSO
- ((pred, succ)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (pred, succ, s_) =>
+ returnm (Some
+ (FENCE_TSO (pred, succ), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s2815_ _s2569_) with
- | Some (s_) => true
- | _ => false
- end)) then
- (match (_s2815_ _s2569_) with
- | Some (s_) =>
- returnm ((Some
- ((FENCEI
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s5220_ _s4974_) with
+ | Some s_ => true
+ | _ => false
+ end then
+ (match (_s5220_ _s4974_) with
+ | Some s_ =>
+ returnm (Some
+ (FENCEI tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s2819_ _s2569_) with
- | Some (s_) => true
- | _ => false
- end)) then
- (match (_s2819_ _s2569_) with
- | Some (s_) =>
- returnm ((Some
- ((ECALL
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s5224_ _s4974_) with
+ | Some s_ => true
+ | _ => false
+ end then
+ (match (_s5224_ _s4974_) with
+ | Some s_ =>
+ returnm (Some
+ (ECALL tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s2823_ _s2569_) with
- | Some (s_) => true
- | _ => false
- end)) then
- (match (_s2823_ _s2569_) with
- | Some (s_) =>
- returnm ((Some
- ((MRET
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s5228_ _s4974_) with
+ | Some s_ => true
+ | _ => false
+ end then
+ (match (_s5228_ _s4974_) with
+ | Some s_ =>
+ returnm (Some
+ (MRET tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s2827_ _s2569_) with
- | Some (s_) => true
- | _ => false
- end)) then
- (match (_s2827_ _s2569_) with
- | Some (s_) =>
- returnm ((Some
- ((SRET
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s5232_ _s4974_) with
+ | Some s_ => true
+ | _ => false
+ end then
+ (match (_s5232_ _s4974_) with
+ | Some s_ =>
+ returnm (Some
+ (SRET tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s2831_ _s2569_) with
- | Some (s_) => true
- | _ => false
- end)) then
- (match (_s2831_ _s2569_) with
- | Some (s_) =>
- returnm ((Some
- ((EBREAK
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s5236_ _s4974_) with
+ | Some s_ => true
+ | _ => false
+ end then
+ (match (_s5236_ _s4974_) with
+ | Some s_ =>
+ returnm (Some
+ (EBREAK tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
- else if ((match (_s2835_ _s2569_) with
- | Some (s_) => true
- | _ => false
- end)) then
- (match (_s2835_ _s2569_) with
- | Some (s_) =>
- returnm ((Some
- ((WFI
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
+ else if match (_s5240_ _s4974_) with
+ | Some s_ => true
+ | _ => false
+ end then
+ (match (_s5240_ _s4974_) with
+ | Some s_ =>
+ returnm (Some
+ (WFI tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2839_ _s2569_) >>= fun w__72 : option ((mword 5 * mword 5 * string)) =>
- (if ((match w__72 with
- | Some ((rs1, rs2, s_)) => true
- | _ => false
- end)) then
- (_s2839_ _s2569_) >>= fun w__73 : option ((mword 5 * mword 5 * string)) =>
+ (_s5244_ _s4974_) >>= fun w__72 : option ((mword 5 * mword 5 * string)) =>
+ (if match w__72 with
+ | Some (rs1, rs2, s_) => true
+ | _ => false
+ end then
+ (_s5244_ _s4974_) >>= fun w__73 : option ((mword 5 * mword 5 * string)) =>
(match w__73 with
- | Some ((rs1, rs2, s_)) =>
- returnm ((Some
- ((SFENCE_VMA
- ((rs1, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >= 0)})))
+ | Some (rs1, rs2, s_) =>
+ returnm (Some
+ (SFENCE_VMA (rs1, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >= 0)})))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)})))
else
- (_s2851_ _s2569_) >>= fun w__76 : option ((word_width * bool * bool * mword 5 * mword 5 * string)) =>
- (if ((match w__76 with
- | Some ((size, aq, rl, rd, rs1, s_)) =>
- true
- | _ => false
- end)) then
- (_s2851_ _s2569_) >>= fun w__77 : option ((word_width * bool * bool * mword 5 * mword 5 * string)) =>
+ (_s5256_ _s4974_) >>= fun w__76 : option ((word_width * bool * bool * mword 5 * mword 5 * string)) =>
+ (if match w__76 with
+ | Some (size, aq, rl, rd, rs1, s_) => true
+ | _ => false
+ end then
+ (_s5256_ _s4974_) >>= fun w__77 : option ((word_width * bool * bool * mword 5 * mword 5 * string)) =>
(match w__77 with
- | Some ((size, aq, rl, rd, rs1, s_)) =>
- returnm ((Some
- ((LOADRES
- ((aq, rl, rs1, size, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ | Some (size, aq, rl, rd, rs1, s_) =>
+ returnm (Some
+ (LOADRES
+ (aq, rl, rs1, size, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s2869_ _s2569_) >>= fun w__80 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
- (if ((match w__80 with
- | Some
- ((size, aq, rl, rd, rs1, rs2, s_)) =>
- true
- | _ => false
- end)) then
- (_s2869_ _s2569_) >>= fun w__81 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5274_ _s4974_) >>= fun w__80 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (if match w__80 with
+ | Some (size, aq, rl, rd, rs1, rs2, s_) =>
+ true
+ | _ => false
+ end then
+ (_s5274_ _s4974_) >>= fun w__81 : option ((word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
(match w__81 with
- | Some ((size, aq, rl, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((STORECON
- ((aq, rl, rs2, rs1, size, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ | Some (size, aq, rl, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (STORECON
+ (aq, rl, rs2, rs1, size, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s2891_ _s2569_) >>= fun w__84 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
- (if ((match w__84 with
- | Some
- ((op, width, aq, rl, rd, rs1, rs2, s_)) =>
- true
- | _ => false
- end)) then
- (_s2891_ _s2569_) >>= fun w__85 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5296_ _s4974_) >>= fun w__84 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (if match w__84 with
+ | Some
+ (op, width, aq, rl, rd, rs2, rs1, s_) =>
+ true
+ | _ => false
+ end then
+ (_s5296_ _s4974_) >>= fun w__85 : option ((amoop * word_width * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
(match w__85 with
| Some
- ((op, width, aq, rl, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((AMO
- ((op, aq, rl, rs2, rs1, width, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (op, width, aq, rl, rd, rs2, rs1, s_) =>
+ returnm (Some
+ (AMO
+ (op, aq, rl, rs2, rs1, width, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
- else if ((match (_s2915_ _s2569_) with
- | Some (s_) => true
- | _ => false
- end)) then
- (match (_s2915_ _s2569_) with
- | Some (s_) =>
- returnm ((Some
- ((C_NOP
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length arg_))
- (projT1
- (string_length s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ else if match (_s5322_ _s4974_) with
+ | Some s_ => true
+ | _ => false
+ end then
+ (match (_s5322_ _s4974_) with
+ | Some s_ =>
+ returnm (Some
+ (C_NOP tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s2919_ _s2569_) >>= fun w__90 : option ((mword 3 * mword 8 * string)) =>
- (if ((match w__90 with
- | Some ((rdc, nzimm, s_)) =>
- neq_vec nzimm
- (vec_of_bits [B0;B0;B0;B0;B0;B0;B0;B0]
- : mword 8)
- | _ => false
- end)) then
- (_s2919_ _s2569_) >>= fun w__91 : option ((mword 3 * mword 8 * string)) =>
+ (_s5326_ _s4974_) >>= fun w__90 : option ((mword 3 * mword 8 * string)) =>
+ (if match w__90 with
+ | Some (rdc, nzimm, s_) =>
+ neq_vec nzimm (Ox"00" : mword 8)
+ | _ => false
+ end then
+ (_s5326_ _s4974_) >>= fun w__91 : option ((mword 3 * mword 8 * string)) =>
(match w__91 with
- | Some ((rdc, nzimm, s_)) =>
- returnm ((Some
- ((C_ADDI4SPN
- ((rdc, nzimm)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ | Some (rdc, nzimm, s_) =>
+ returnm (Some
+ (C_ADDI4SPN
+ (rdc, nzimm), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s2931_ _s2569_) >>= fun w__94 : option ((mword 3 * mword 3 * mword 5 * string)) =>
- (if ((match w__94 with
- | Some ((rdc, rsc, uimm, s_)) =>
- true
- | _ => false
- end)) then
- (_s2931_ _s2569_) >>= fun w__95 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (_s5338_ _s4974_) >>= fun w__94 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (if match w__94 with
+ | Some (rdc, rsc, uimm, s_) =>
+ true
+ | _ => false
+ end then
+ (_s5338_ _s4974_) >>= fun w__95 : option ((mword 3 * mword 3 * mword 5 * string)) =>
(match w__95 with
- | Some ((rdc, rsc, uimm, s_)) =>
- returnm ((Some
- ((C_LW
- ((uimm, rsc, rdc)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ | Some (rdc, rsc, uimm, s_) =>
+ returnm (Some
+ (C_LW
+ (uimm, rsc, rdc), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s2947_ _s2569_) >>= fun w__98 : option ((mword 3 * mword 3 * mword 5 * string)) =>
- (if ((match w__98 with
- | Some
- ((rdc, rsc, uimm, s_)) =>
- Z.eqb 64 64
- | _ => false
- end)) then
- (_s2947_ _s2569_) >>= fun w__99 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (_s5354_ _s4974_) >>= fun w__98 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (if match w__98 with
+ | Some (rdc, rsc, uimm, s_) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s5354_ _s4974_) >>= fun w__99 : option ((mword 3 * mword 3 * mword 5 * string)) =>
(match w__99 with
- | Some ((rdc, rsc, uimm, s_)) =>
- returnm ((Some
- ((C_LD
- ((uimm, rsc, rdc)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ | Some (rdc, rsc, uimm, s_) =>
+ returnm (Some
+ (C_LD
+ (uimm, rsc, rdc), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s2963_ _s2569_) >>= fun w__102 : option ((mword 3 * mword 3 * mword 5 * string)) =>
- (if ((match w__102 with
- | Some
- ((rsc1, rsc2, uimm, s_)) =>
- true
- | _ => false
- end)) then
- (_s2963_ _s2569_) >>= fun w__103 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (_s5370_ _s4974_) >>= fun w__102 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (if match w__102 with
+ | Some
+ (rsc1, rsc2, uimm, s_) =>
+ true
+ | _ => false
+ end then
+ (_s5370_ _s4974_) >>= fun w__103 : option ((mword 3 * mword 3 * mword 5 * string)) =>
(match w__103 with
| Some
- ((rsc1, rsc2, uimm, s_)) =>
- returnm ((Some
- ((C_SW
- ((uimm, rsc1, rsc2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsc1, rsc2, uimm, s_) =>
+ returnm (Some
+ (C_SW
+ (uimm, rsc1, rsc2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s2979_ _s2569_) >>= fun w__106 : option ((mword 3 * mword 3 * mword 5 * string)) =>
- (if ((match w__106 with
- | Some
- ((rsc1, rsc2, uimm, s_)) =>
- Z.eqb 64 64
- | _ => false
- end)) then
- (_s2979_ _s2569_) >>= fun w__107 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (_s5386_ _s4974_) >>= fun w__106 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (if match w__106 with
+ | Some
+ (rsc1, rsc2, uimm, s_) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s5386_ _s4974_) >>= fun w__107 : option ((mword 3 * mword 3 * mword 5 * string)) =>
(match w__107 with
| Some
- ((rsc1, rsc2, uimm, s_)) =>
- returnm ((Some
- ((C_SD
- ((uimm, rsc1, rsc2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsc1, rsc2, uimm, s_) =>
+ returnm (Some
+ (C_SD
+ (uimm, rsc1, rsc2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s2995_ _s2569_) >>= fun w__110 : option ((mword 5 * mword 6 * string)) =>
- (if ((match w__110 with
- | Some
- ((rsd, nzi, s_)) =>
- andb
- (neq_vec nzi
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ => false
- end)) then
- (_s2995_ _s2569_) >>= fun w__111 : option ((mword 5 * mword 6 * string)) =>
+ (_s5402_ _s4974_) >>= fun w__110 : option ((mword 5 * mword 6 * string)) =>
+ (if match w__110 with
+ | Some (rsd, nzi, s_) =>
+ andb
+ (neq_vec nzi
+ ('b"000000"
+ : mword 6))
+ (neq_vec rsd
+ zreg)
+ | _ => false
+ end then
+ (_s5402_ _s4974_) >>= fun w__111 : option ((mword 5 * mword 6 * string)) =>
(match w__111 with
- | Some
- ((rsd, nzi, s_)) =>
- returnm ((Some
- ((C_ADDI
- ((nzi, rsd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ | Some (rsd, nzi, s_) =>
+ returnm (Some
+ (C_ADDI
+ (nzi, rsd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3007_ _s2569_) >>= fun w__114 : option ((mword 11 * string)) =>
- (if ((match w__114 with
- | Some
- ((imm, s_)) =>
- Z.eqb 64 32
- | _ => false
- end)) then
- (_s3007_ _s2569_) >>= fun w__115 : option ((mword 11 * string)) =>
+ (_s5414_ _s4974_) >>= fun w__114 : option ((mword 11 * string)) =>
+ (if match w__114 with
+ | Some (imm, s_) =>
+ Z.eqb 64 32
+ | _ => false
+ end then
+ (_s5414_ _s4974_) >>= fun w__115 : option ((mword 11 * string)) =>
(match w__115 with
- | Some ((imm, s_)) =>
- returnm ((Some
- ((C_JAL
- (imm), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ | Some (imm, s_) =>
+ returnm (Some
+ (C_JAL
+ imm, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3015_ _s2569_) >>= fun w__118 : option ((mword 5 * mword 6 * string)) =>
- (if ((match w__118 with
- | Some
- ((rsd, imm, s_)) =>
- Z.eqb 64
- 64
- | _ => false
- end)) then
- (_s3015_ _s2569_) >>= fun w__119 : option ((mword 5 * mword 6 * string)) =>
+ (_s5422_ _s4974_) >>= fun w__118 : option ((mword 5 * mword 6 * string)) =>
+ (if match w__118 with
+ | Some
+ (rsd, imm, s_) =>
+ Z.eqb 64 64
+ | _ => false
+ end then
+ (_s5422_ _s4974_) >>= fun w__119 : option ((mword 5 * mword 6 * string)) =>
(match w__119 with
| Some
- ((rsd, imm, s_)) =>
- returnm ((Some
- ((C_ADDIW
- ((imm, rsd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, imm, s_) =>
+ returnm (Some
+ (C_ADDIW
+ (imm, rsd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3027_ _s2569_) >>= fun w__122 : option ((mword 5 * mword 6 * string)) =>
- (if ((match w__122 with
- | Some
- ((rd, imm, s_)) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end)) then
- (_s3027_
- _s2569_) >>= fun w__123 : option ((mword 5 * mword 6 * string)) =>
+ (_s5434_ _s4974_) >>= fun w__122 : option ((mword 5 * mword 6 * string)) =>
+ (if match w__122 with
+ | Some
+ (rd, imm, s_) =>
+ neq_vec
+ rd zreg
+ | _ => false
+ end then
+ (_s5434_
+ _s4974_) >>= fun w__123 : option ((mword 5 * mword 6 * string)) =>
(match w__123 with
| Some
- ((rd, imm, s_)) =>
- returnm ((Some
- ((C_LI
- ((imm, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rd, imm, s_) =>
+ returnm (Some
+ (C_LI
+ (imm, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3039_
- _s2569_) >>= fun w__126 : option ((mword 6 * string)) =>
- (if ((match w__126 with
- | Some
- ((imm, s_)) =>
- neq_vec
- imm
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
- then
- (_s3039_
- _s2569_) >>= fun w__127 : option ((mword 6 * string)) =>
+ (_s5446_
+ _s4974_) >>= fun w__126 : option ((mword 6 * string)) =>
+ (if match w__126 with
+ | Some
+ (imm, s_) =>
+ neq_vec
+ imm
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end then
+ (_s5446_
+ _s4974_) >>= fun w__127 : option ((mword 6 * string)) =>
(match w__127 with
| Some
- ((imm, s_)) =>
- returnm ((Some
- ((C_ADDI16SP
- (imm), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (imm, s_) =>
+ returnm (Some
+ (C_ADDI16SP
+ imm, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3047_
- _s2569_) >>= fun w__130 : option ((mword 5 * mword 6 * string)) =>
- (if ((match w__130 with
- | Some
- ((rd, imm, s_)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- sp))))
- (neq_vec
- imm
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)))
- : bool)
- | _ =>
- false
- end))
+ (_s5454_
+ _s4974_) >>= fun w__130 : option ((mword 5 * mword 6 * string)) =>
+ (if match w__130 with
+ | Some
+ (rd, imm, s_) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (andb
+ (neq_vec
+ rd
+ sp)
+ (neq_vec
+ imm
+ ('b"000000"
+ : mword 6)))
+ | _ =>
+ false
+ end
then
- (_s3047_
- _s2569_) >>= fun w__131 : option ((mword 5 * mword 6 * string)) =>
+ (_s5454_
+ _s4974_) >>= fun w__131 : option ((mword 5 * mword 6 * string)) =>
(match w__131 with
| Some
- ((rd, imm, s_)) =>
- returnm ((Some
- ((C_LUI
- ((imm, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rd, imm, s_) =>
+ returnm (Some
+ (C_LUI
+ (imm, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3059_
- _s2569_) >>= fun w__134 : option ((mword 3 * mword 6 * string)) =>
- (if ((match w__134 with
- | Some
- ((rsd, shamt, s_)) =>
- neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
+ (_s5466_
+ _s4974_) >>= fun w__134 : option ((mword 3 * mword 6 * string)) =>
+ (if match w__134 with
+ | Some
+ (rsd, shamt, s_) =>
+ neq_vec
+ shamt
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end
then
- (_s3059_
- _s2569_) >>= fun w__135 : option ((mword 3 * mword 6 * string)) =>
+ (_s5466_
+ _s4974_) >>= fun w__135 : option ((mword 3 * mword 6 * string)) =>
(match w__135 with
| Some
- ((rsd, shamt, s_)) =>
- returnm ((Some
- ((C_SRLI
- ((shamt, rsd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, shamt, s_) =>
+ returnm (Some
+ (C_SRLI
+ (shamt, rsd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3071_
- _s2569_) >>= fun w__138 : option ((mword 3 * mword 6 * string)) =>
- (if ((match w__138 with
- | Some
- ((rsd, shamt, s_)) =>
- neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6)
- | _ =>
- false
- end))
+ (_s5478_
+ _s4974_) >>= fun w__138 : option ((mword 3 * mword 6 * string)) =>
+ (if match w__138 with
+ | Some
+ (rsd, shamt, s_) =>
+ neq_vec
+ shamt
+ ('b"000000"
+ : mword 6)
+ | _ =>
+ false
+ end
then
- (_s3071_
- _s2569_) >>= fun w__139 : option ((mword 3 * mword 6 * string)) =>
+ (_s5478_
+ _s4974_) >>= fun w__139 : option ((mword 3 * mword 6 * string)) =>
(match w__139 with
| Some
- ((rsd, shamt, s_)) =>
- returnm ((Some
- ((C_SRAI
- ((shamt, rsd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, shamt, s_) =>
+ returnm (Some
+ (C_SRAI
+ (shamt, rsd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3083_
- _s2569_) >>= fun w__142 : option ((mword 3 * mword 6 * string)) =>
+ (_s5490_
+ _s4974_) >>= fun w__142 : option ((mword 3 * mword 6 * string)) =>
(if
- ((match w__142 with
- | Some
- ((rsd, imm, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__142 with
+ | Some
+ (rsd, imm, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3083_
- _s2569_) >>= fun w__143 : option ((mword 3 * mword 6 * string)) =>
+ (_s5490_
+ _s4974_) >>= fun w__143 : option ((mword 3 * mword 6 * string)) =>
(match w__143 with
| Some
- ((rsd, imm, s_)) =>
- returnm ((Some
- ((C_ANDI
- ((imm, rsd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, imm, s_) =>
+ returnm (Some
+ (C_ANDI
+ (imm, rsd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3095_
- _s2569_) >>= fun w__146 : option ((mword 3 * mword 3 * string)) =>
+ (_s5502_
+ _s4974_) >>= fun w__146 : option ((mword 3 * mword 3 * string)) =>
(if
- ((match w__146 with
- | Some
- ((rsd, rs2, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__146 with
+ | Some
+ (rsd, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3095_
- _s2569_) >>= fun w__147 : option ((mword 3 * mword 3 * string)) =>
+ (_s5502_
+ _s4974_) >>= fun w__147 : option ((mword 3 * mword 3 * string)) =>
(match w__147 with
| Some
- ((rsd, rs2, s_)) =>
- returnm ((Some
- ((C_SUB
- ((rsd, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, rs2, s_) =>
+ returnm (Some
+ (C_SUB
+ (rsd, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3107_
- _s2569_) >>= fun w__150 : option ((mword 3 * mword 3 * string)) =>
+ (_s5514_
+ _s4974_) >>= fun w__150 : option ((mword 3 * mword 3 * string)) =>
(if
- ((match w__150 with
- | Some
- ((rsd, rs2, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__150 with
+ | Some
+ (rsd, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3107_
- _s2569_) >>= fun w__151 : option ((mword 3 * mword 3 * string)) =>
+ (_s5514_
+ _s4974_) >>= fun w__151 : option ((mword 3 * mword 3 * string)) =>
(match w__151 with
| Some
- ((rsd, rs2, s_)) =>
- returnm ((Some
- ((C_XOR
- ((rsd, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, rs2, s_) =>
+ returnm (Some
+ (C_XOR
+ (rsd, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3119_
- _s2569_) >>= fun w__154 : option ((mword 3 * mword 3 * string)) =>
+ (_s5526_
+ _s4974_) >>= fun w__154 : option ((mword 3 * mword 3 * string)) =>
(if
- ((match w__154 with
- | Some
- ((rsd, rs2, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__154 with
+ | Some
+ (rsd, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3119_
- _s2569_) >>= fun w__155 : option ((mword 3 * mword 3 * string)) =>
+ (_s5526_
+ _s4974_) >>= fun w__155 : option ((mword 3 * mword 3 * string)) =>
(match w__155 with
| Some
- ((rsd, rs2, s_)) =>
- returnm ((Some
- ((C_OR
- ((rsd, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, rs2, s_) =>
+ returnm (Some
+ (C_OR
+ (rsd, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3131_
- _s2569_) >>= fun w__158 : option ((mword 3 * mword 3 * string)) =>
+ (_s5538_
+ _s4974_) >>= fun w__158 : option ((mword 3 * mword 3 * string)) =>
(if
- ((match w__158 with
- | Some
- ((rsd, rs2, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__158 with
+ | Some
+ (rsd, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3131_
- _s2569_) >>= fun w__159 : option ((mword 3 * mword 3 * string)) =>
+ (_s5538_
+ _s4974_) >>= fun w__159 : option ((mword 3 * mword 3 * string)) =>
(match w__159 with
| Some
- ((rsd, rs2, s_)) =>
- returnm ((Some
- ((C_AND
- ((rsd, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, rs2, s_) =>
+ returnm (Some
+ (C_AND
+ (rsd, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3143_
- _s2569_) >>= fun w__162 : option ((mword 3 * mword 3 * string)) =>
+ (_s5550_
+ _s4974_) >>= fun w__162 : option ((mword 3 * mword 3 * string)) =>
(if
- ((match w__162 with
- | Some
- ((rsd, rs2, s_)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__162 with
+ | Some
+ (rsd, rs2, s_) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s3143_
- _s2569_) >>= fun w__163 : option ((mword 3 * mword 3 * string)) =>
+ (_s5550_
+ _s4974_) >>= fun w__163 : option ((mword 3 * mword 3 * string)) =>
(match w__163 with
| Some
- ((rsd, rs2, s_)) =>
- returnm ((Some
- ((C_SUBW
- ((rsd, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, rs2, s_) =>
+ returnm (Some
+ (C_SUBW
+ (rsd, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3155_
- _s2569_) >>= fun w__166 : option ((mword 3 * mword 3 * string)) =>
+ (_s5562_
+ _s4974_) >>= fun w__166 : option ((mword 3 * mword 3 * string)) =>
(if
- ((match w__166 with
- | Some
- ((rsd, rs2, s_)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__166 with
+ | Some
+ (rsd, rs2, s_) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s3155_
- _s2569_) >>= fun w__167 : option ((mword 3 * mword 3 * string)) =>
+ (_s5562_
+ _s4974_) >>= fun w__167 : option ((mword 3 * mword 3 * string)) =>
(match w__167 with
| Some
- ((rsd, rs2, s_)) =>
- returnm ((Some
- ((C_ADDW
- ((rsd, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, rs2, s_) =>
+ returnm (Some
+ (C_ADDW
+ (rsd, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3167_
- _s2569_) >>= fun w__170 : option ((mword 11 * string)) =>
+ (_s5574_
+ _s4974_) >>= fun w__170 : option ((mword 11 * string)) =>
(if
- ((match w__170 with
- | Some
- ((imm, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__170 with
+ | Some
+ (imm, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3167_
- _s2569_) >>= fun w__171 : option ((mword 11 * string)) =>
+ (_s5574_
+ _s4974_) >>= fun w__171 : option ((mword 11 * string)) =>
(match w__171 with
| Some
- ((imm, s_)) =>
- returnm ((Some
- ((C_J
- (imm), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (imm, s_) =>
+ returnm (Some
+ (C_J
+ imm, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3175_
- _s2569_) >>= fun w__174 : option ((mword 3 * mword 8 * string)) =>
+ (_s5582_
+ _s4974_) >>= fun w__174 : option ((mword 3 * mword 8 * string)) =>
(if
- ((match w__174 with
- | Some
- ((rs, imm, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__174 with
+ | Some
+ (rs, imm, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3175_
- _s2569_) >>= fun w__175 : option ((mword 3 * mword 8 * string)) =>
+ (_s5582_
+ _s4974_) >>= fun w__175 : option ((mword 3 * mword 8 * string)) =>
(match w__175 with
| Some
- ((rs, imm, s_)) =>
- returnm ((Some
- ((C_BEQZ
- ((imm, rs)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rs, imm, s_) =>
+ returnm (Some
+ (C_BEQZ
+ (imm, rs), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3187_
- _s2569_) >>= fun w__178 : option ((mword 3 * mword 8 * string)) =>
+ (_s5594_
+ _s4974_) >>= fun w__178 : option ((mword 3 * mword 8 * string)) =>
(if
- ((match w__178 with
- | Some
- ((rs, imm, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__178 with
+ | Some
+ (rs, imm, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3187_
- _s2569_) >>= fun w__179 : option ((mword 3 * mword 8 * string)) =>
+ (_s5594_
+ _s4974_) >>= fun w__179 : option ((mword 3 * mword 8 * string)) =>
(match w__179 with
| Some
- ((rs, imm, s_)) =>
- returnm ((Some
- ((C_BNEZ
- ((imm, rs)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rs, imm, s_) =>
+ returnm (Some
+ (C_BNEZ
+ (imm, rs), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3199_
- _s2569_) >>= fun w__182 : option ((mword 5 * mword 6 * string)) =>
+ (_s5606_
+ _s4974_) >>= fun w__182 : option ((mword 5 * mword 6 * string)) =>
(if
- ((match w__182 with
- | Some
- ((rsd, shamt, s_)) =>
- andb
- (neq_vec
- shamt
- (vec_of_bits [B0;B0;B0;B0;B0;B0]
- : mword 6))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__182 with
+ | Some
+ (rsd, shamt, s_) =>
+ andb
+ (neq_vec
+ shamt
+ ('b"000000"
+ : mword 6))
+ (neq_vec
+ rsd
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s3199_
- _s2569_) >>= fun w__183 : option ((mword 5 * mword 6 * string)) =>
+ (_s5606_
+ _s4974_) >>= fun w__183 : option ((mword 5 * mword 6 * string)) =>
(match w__183 with
| Some
- ((rsd, shamt, s_)) =>
- returnm ((Some
- ((C_SLLI
- ((shamt, rsd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, shamt, s_) =>
+ returnm (Some
+ (C_SLLI
+ (shamt, rsd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3211_
- _s2569_) >>= fun w__186 : option ((mword 5 * mword 6 * string)) =>
+ (_s5618_
+ _s4974_) >>= fun w__186 : option ((mword 5 * mword 6 * string)) =>
(if
- ((match w__186 with
- | Some
- ((rd, uimm, s_)) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__186 with
+ | Some
+ (rd, uimm, s_) =>
+ neq_vec
+ rd
+ zreg
+ | _ =>
+ false
+ end
then
- (_s3211_
- _s2569_) >>= fun w__187 : option ((mword 5 * mword 6 * string)) =>
+ (_s5618_
+ _s4974_) >>= fun w__187 : option ((mword 5 * mword 6 * string)) =>
(match w__187 with
| Some
- ((rd, uimm, s_)) =>
- returnm ((Some
- ((C_LWSP
- ((uimm, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rd, uimm, s_) =>
+ returnm (Some
+ (C_LWSP
+ (uimm, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3223_
- _s2569_) >>= fun w__190 : option ((mword 5 * mword 6 * string)) =>
+ (_s5630_
+ _s4974_) >>= fun w__190 : option ((mword 5 * mword 6 * string)) =>
(if
- ((match w__190 with
- | Some
- ((rd, uimm, s_)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- (Z.eqb
- 64
- 64)
- | _ =>
- false
- end))
+ match w__190 with
+ | Some
+ (rd, uimm, s_) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
then
- (_s3223_
- _s2569_) >>= fun w__191 : option ((mword 5 * mword 6 * string)) =>
+ (_s5630_
+ _s4974_) >>= fun w__191 : option ((mword 5 * mword 6 * string)) =>
(match w__191 with
| Some
- ((rd, uimm, s_)) =>
- returnm ((Some
- ((C_LDSP
- ((uimm, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rd, uimm, s_) =>
+ returnm (Some
+ (C_LDSP
+ (uimm, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3235_
- _s2569_) >>= fun w__194 : option ((mword 5 * mword 6 * string)) =>
+ (_s5642_
+ _s4974_) >>= fun w__194 : option ((mword 5 * mword 6 * string)) =>
(if
- ((match w__194 with
- | Some
- ((rd, uimm, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__194 with
+ | Some
+ (rd, uimm, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3235_
- _s2569_) >>= fun w__195 : option ((mword 5 * mword 6 * string)) =>
+ (_s5642_
+ _s4974_) >>= fun w__195 : option ((mword 5 * mword 6 * string)) =>
(match w__195 with
| Some
- ((rd, uimm, s_)) =>
- returnm ((Some
- ((C_SWSP
- ((uimm, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rd, uimm, s_) =>
+ returnm (Some
+ (C_SWSP
+ (uimm, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3247_
- _s2569_) >>= fun w__198 : option ((mword 5 * mword 6 * string)) =>
+ (_s5654_
+ _s4974_) >>= fun w__198 : option ((mword 5 * mword 6 * string)) =>
(if
- ((match w__198 with
- | Some
- ((rs2, uimm, s_)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__198 with
+ | Some
+ (rs2, uimm, s_) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s3247_
- _s2569_) >>= fun w__199 : option ((mword 5 * mword 6 * string)) =>
+ (_s5654_
+ _s4974_) >>= fun w__199 : option ((mword 5 * mword 6 * string)) =>
(match w__199 with
| Some
- ((rs2, uimm, s_)) =>
- returnm ((Some
- ((C_SDSP
- ((uimm, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rs2, uimm, s_) =>
+ returnm (Some
+ (C_SDSP
+ (uimm, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3259_
- _s2569_) >>= fun w__202 : option ((mword 5 * string)) =>
+ (_s5666_
+ _s4974_) >>= fun w__202 : option ((mword 5 * string)) =>
(if
- ((match w__202 with
- | Some
- ((rs1, s_)) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs1))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__202 with
+ | Some
+ (rs1, s_) =>
+ neq_vec
+ rs1
+ zreg
+ | _ =>
+ false
+ end
then
- (_s3259_
- _s2569_) >>= fun w__203 : option ((mword 5 * string)) =>
+ (_s5666_
+ _s4974_) >>= fun w__203 : option ((mword 5 * string)) =>
(match w__203 with
| Some
- ((rs1, s_)) =>
- returnm ((Some
- ((C_JR
- (rs1), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rs1, s_) =>
+ returnm (Some
+ (C_JR
+ rs1, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3267_
- _s2569_) >>= fun w__206 : option ((mword 5 * string)) =>
+ (_s5674_
+ _s4974_) >>= fun w__206 : option ((mword 5 * string)) =>
(if
- ((match w__206 with
- | Some
- ((rs1, s_)) =>
- projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs1))
- (projT1
- (regidx_to_regno
- zreg)))
- | _ =>
- false
- end))
+ match w__206 with
+ | Some
+ (rs1, s_) =>
+ neq_vec
+ rs1
+ zreg
+ | _ =>
+ false
+ end
then
- (_s3267_
- _s2569_) >>= fun w__207 : option ((mword 5 * string)) =>
+ (_s5674_
+ _s4974_) >>= fun w__207 : option ((mword 5 * string)) =>
(match w__207 with
| Some
- ((rs1, s_)) =>
- returnm ((Some
- ((C_JALR
- (rs1), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rs1, s_) =>
+ returnm (Some
+ (C_JALR
+ rs1, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3275_
- _s2569_) >>= fun w__210 : option ((mword 5 * mword 5 * string)) =>
+ (_s5682_
+ _s4974_) >>= fun w__210 : option ((mword 5 * mword 5 * string)) =>
(if
- ((match w__210 with
- | Some
- ((rd, rs2, s_)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs2))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__210 with
+ | Some
+ (rd, rs2, s_) =>
+ andb
+ (neq_vec
+ rd
+ zreg)
+ (neq_vec
+ rs2
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s3275_
- _s2569_) >>= fun w__211 : option ((mword 5 * mword 5 * string)) =>
+ (_s5682_
+ _s4974_) >>= fun w__211 : option ((mword 5 * mword 5 * string)) =>
(match w__211 with
| Some
- ((rd, rs2, s_)) =>
- returnm ((Some
- ((C_MV
- ((rd, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rd, rs2, s_) =>
+ returnm (Some
+ (C_MV
+ (rd, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else if
- ((match (_s3287_
- _s2569_) with
- | Some
- (s_) =>
- true
- | _ =>
- false
- end))
+ match (_s5694_
+ _s4974_) with
+ | Some
+ s_ =>
+ true
+ | _ =>
+ false
+ end
then
- (match (_s3287_
- _s2569_) with
+ (match (_s5694_
+ _s4974_) with
| Some
- (s_) =>
- returnm ((Some
- ((C_EBREAK
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ s_ =>
+ returnm (Some
+ (C_EBREAK
+ tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3291_
- _s2569_) >>= fun w__216 : option ((mword 5 * mword 5 * string)) =>
+ (_s5698_
+ _s4974_) >>= fun w__216 : option ((mword 5 * mword 5 * string)) =>
(if
- ((match w__216 with
- | Some
- ((rsd, rs2, s_)) =>
- andb
- (projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rsd))
- (projT1
- (regidx_to_regno
- zreg))))
- ((projT1
- (neq_int
- (projT1
- (regidx_to_regno
- rs2))
- (projT1
- (regidx_to_regno
- zreg))))
- : bool)
- | _ =>
- false
- end))
+ match w__216 with
+ | Some
+ (rsd, rs2, s_) =>
+ andb
+ (neq_vec
+ rsd
+ zreg)
+ (neq_vec
+ rs2
+ zreg)
+ | _ =>
+ false
+ end
then
- (_s3291_
- _s2569_) >>= fun w__217 : option ((mword 5 * mword 5 * string)) =>
+ (_s5698_
+ _s4974_) >>= fun w__217 : option ((mword 5 * mword 5 * string)) =>
(match w__217 with
| Some
- ((rsd, rs2, s_)) =>
- returnm ((Some
- ((C_ADD
- ((rsd, rs2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rsd, rs2, s_) =>
+ returnm (Some
+ (C_ADD
+ (rsd, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3303_
- _s2569_) >>= fun w__220 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5710_
+ _s4974_) >>= fun w__220 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
(if
- ((match w__220 with
- | Some
- ((high, signed1, signed2, rd, rs1, rs2, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__220 with
+ | Some
+ (high, signed1, signed2, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3303_
- _s2569_) >>= fun w__221 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5710_
+ _s4974_) >>= fun w__221 : option ((bool * bool * bool * mword 5 * mword 5 * mword 5 * string)) =>
(match w__221 with
| Some
- ((high, signed1, signed2, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((MUL
- ((rs2, rs1, rd, high, signed1, signed2)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (high, signed1, signed2, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (MUL
+ (rs2, rs1, rd, high, signed1, signed2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3320_
- _s2569_) >>= fun w__224 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5727_
+ _s4974_) >>= fun w__224 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
(if
- ((match w__224 with
- | Some
- ((s, rd, rs1, rs2, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__224 with
+ | Some
+ (s, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3320_
- _s2569_) >>= fun w__225 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5727_
+ _s4974_) >>= fun w__225 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
(match w__225 with
| Some
- ((s, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((DIV
- ((rs2, rs1, rd, s)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (s, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (DIV
+ (rs2, rs1, rd, s), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3338_
- _s2569_) >>= fun w__228 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5745_
+ _s4974_) >>= fun w__228 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
(if
- ((match w__228 with
- | Some
- ((s, rd, rs1, rs2, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__228 with
+ | Some
+ (s, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3338_
- _s2569_) >>= fun w__229 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5745_
+ _s4974_) >>= fun w__229 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
(match w__229 with
| Some
- ((s, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((REM
- ((rs2, rs1, rd, s)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (s, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (REM
+ (rs2, rs1, rd, s), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3356_
- _s2569_) >>= fun w__232 : option ((mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5763_
+ _s4974_) >>= fun w__232 : option ((mword 5 * mword 5 * mword 5 * string)) =>
(if
- ((match w__232 with
- | Some
- ((rd, rs1, rs2, s_)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__232 with
+ | Some
+ (rd, rs1, rs2, s_) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s3356_
- _s2569_) >>= fun w__233 : option ((mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5763_
+ _s4974_) >>= fun w__233 : option ((mword 5 * mword 5 * mword 5 * string)) =>
(match w__233 with
| Some
- ((rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((MULW
- ((rs2, rs1, rd)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (rd, rs1, rs2, s_) =>
+ returnm (Some
+ (MULW
+ (rs2, rs1, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3372_
- _s2569_) >>= fun w__236 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5779_
+ _s4974_) >>= fun w__236 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
(if
- ((match w__236 with
- | Some
- ((s, rd, rs1, rs2, s_)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__236 with
+ | Some
+ (s, rd, rs1, rs2, s_) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s3372_
- _s2569_) >>= fun w__237 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5779_
+ _s4974_) >>= fun w__237 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
(match w__237 with
| Some
- ((s, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((DIVW
- ((rs2, rs1, rd, s)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (s, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (DIVW
+ (rs2, rs1, rd, s), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3391_
- _s2569_) >>= fun w__240 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5798_
+ _s4974_) >>= fun w__240 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
(if
- ((match w__240 with
- | Some
- ((s, rd, rs1, rs2, s_)) =>
- Z.eqb
- 64
- 64
- | _ =>
- false
- end))
+ match w__240 with
+ | Some
+ (s, rd, rs1, rs2, s_) =>
+ Z.eqb
+ 64
+ 64
+ | _ =>
+ false
+ end
then
- (_s3391_
- _s2569_) >>= fun w__241 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
+ (_s5798_
+ _s4974_) >>= fun w__241 : option ((bool * mword 5 * mword 5 * mword 5 * string)) =>
(match w__241 with
| Some
- ((s, rd, rs1, rs2, s_)) =>
- returnm ((Some
- ((REMW
- ((rs2, rs1, rd, s)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (s, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (REMW
+ (rs2, rs1, rd, s), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3410_
- _s2569_) >>= fun w__244 : option ((csrop * mword 5 * mword 5 * mword 12 * string)) =>
+ (_s5817_
+ _s4974_) >>= fun w__244 : option ((csrop * mword 5 * mword 12 * mword 5 * string)) =>
(if
- ((match w__244 with
- | Some
- ((op, rd, rs1, csr, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__244 with
+ | Some
+ (op, rd, csr, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3410_
- _s2569_) >>= fun w__245 : option ((csrop * mword 5 * mword 5 * mword 12 * string)) =>
+ (_s5817_
+ _s4974_) >>= fun w__245 : option ((csrop * mword 5 * mword 12 * mword 5 * string)) =>
(match w__245 with
| Some
- ((op, rd, rs1, csr, s_)) =>
- returnm ((Some
- ((CSR
- ((csr, rs1, rd, true, op)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (op, rd, csr, rs1, s_) =>
+ returnm (Some
+ (CSR
+ (csr, rs1, rd, true, op), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3428_
- _s2569_) >>= fun w__248 : option ((csrop * mword 5 * mword 5 * mword 12 * string)) =>
+ (_s5835_
+ _s4974_) >>= fun w__248 : option ((csrop * mword 5 * mword 12 * mword 5 * string)) =>
(if
- ((match w__248 with
- | Some
- ((op, rd, rs1, csr, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__248 with
+ | Some
+ (op, rd, csr, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3428_
- _s2569_) >>= fun w__249 : option ((csrop * mword 5 * mword 5 * mword 12 * string)) =>
+ (_s5835_
+ _s4974_) >>= fun w__249 : option ((csrop * mword 5 * mword 12 * mword 5 * string)) =>
(match w__249 with
| Some
- ((op, rd, rs1, csr, s_)) =>
- returnm ((Some
- ((CSR
- ((csr, rs1, rd, false, op)), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (op, rd, csr, rs1, s_) =>
+ returnm (Some
+ (CSR
+ (csr, rs1, rd, false, op), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else if
- ((match (_s3445_
- _s2569_) with
- | Some
- (s_) =>
- true
- | _ =>
- false
- end))
+ match (_s5852_
+ _s4974_) with
+ | Some
+ s_ =>
+ true
+ | _ =>
+ false
+ end
then
- (match (_s3445_
- _s2569_) with
+ (match (_s5852_
+ _s4974_) with
| Some
- (s_) =>
- returnm ((Some
- ((URET
- (tt), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ s_ =>
+ returnm (Some
+ (URET
+ tt, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3449_
- _s2569_) >>= fun w__254 : option ((mword 32 * string)) =>
+ (_s5856_
+ _s4974_) >>= fun w__254 : option ((word_width * mword 5 * mword 12 * mword 5 * string)) =>
(if
- ((match w__254 with
- | Some
- ((s, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__254 with
+ | Some
+ (width, rd, imm, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3449_
- _s2569_) >>= fun w__255 : option ((mword 32 * string)) =>
+ (_s5856_
+ _s4974_) >>= fun w__255 : option ((word_width * mword 5 * mword 12 * mword 5 * string)) =>
(match w__255 with
| Some
- ((s, s_)) =>
- returnm ((Some
- ((ILLEGAL
- (s), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (width, rd, imm, rs1, s_) =>
+ returnm (Some
+ (LOAD_FP
+ (imm, rs1, rd, width), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- (_s3457_
- _s2569_) >>= fun w__258 : option ((mword 16 * string)) =>
+ (_s5880_
+ _s4974_) >>= fun w__258 : option ((word_width * mword 5 * mword 12 * mword 5 * string)) =>
(if
- ((match w__258 with
- | Some
- ((s, s_)) =>
- true
- | _ =>
- false
- end))
+ match w__258 with
+ | Some
+ (width, rs2, imm, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
then
- (_s3457_
- _s2569_) >>= fun w__259 : option ((mword 16 * string)) =>
+ (_s5880_
+ _s4974_) >>= fun w__259 : option ((word_width * mword 5 * mword 12 * mword 5 * string)) =>
(match w__259 with
| Some
- ((s, s_)) =>
- returnm ((Some
- ((C_ILLEGAL
- (s), build_ex
- (projT1
- (sub_nat
- (projT1
- (string_length
- arg_))
- (projT1
- (string_length
- s_)))))))
- : option ((ast * {n : Z & ArithFact (n >=
- 0)})))
+ (width, rs2, imm, rs1, s_) =>
+ returnm (Some
+ (STORE_FP
+ (imm, rs2, rs1, width), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
| _ =>
exit tt
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
end)
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)})))
else
- returnm (None
- : option ((ast * {n : Z & ArithFact (n >=
+ (_s5904_
+ _s4974_) >>= fun w__262 : option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__262 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s5904_
+ _s4974_) >>= fun w__263 : option ((f_madd_op_S * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__263 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm, s_) =>
+ returnm (Some
+ (F_MADD_TYPE_S
+ (rs3, rs2, rs1, rm, rd, op), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s5929_
+ _s4974_) >>= fun w__266 : option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__266 with
+ | Some
+ (op, rd, rs1, rs2, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s5929_
+ _s4974_) >>= fun w__267 : option ((f_bin_rm_op_S * mword 5 * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__267 with
+ | Some
+ (op, rd, rs1, rs2, rm, s_) =>
+ returnm (Some
+ (F_BIN_RM_TYPE_S
+ (rs2, rs1, rm, rd, op), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s5950_
+ _s4974_) >>= fun w__270 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__270 with
+ | Some
+ (FSQRT_S, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s5950_
+ _s4974_) >>= fun w__271 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__271 with
+ | Some
+ (FSQRT_S, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FSQRT_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s5967_
+ _s4974_) >>= fun w__274 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__274 with
+ | Some
+ (FCVT_W_S, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s5967_
+ _s4974_) >>= fun w__275 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__275 with
+ | Some
+ (FCVT_W_S, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_W_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s5984_
+ _s4974_) >>= fun w__278 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__278 with
+ | Some
+ (FCVT_WU_S, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s5984_
+ _s4974_) >>= fun w__279 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__279 with
+ | Some
+ (FCVT_WU_S, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_WU_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6001_
+ _s4974_) >>= fun w__282 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__282 with
+ | Some
+ (FCVT_S_W, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6001_
+ _s4974_) >>= fun w__283 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__283 with
+ | Some
+ (FCVT_S_W, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_W), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6018_
+ _s4974_) >>= fun w__286 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__286 with
+ | Some
+ (FCVT_S_WU, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6018_
+ _s4974_) >>= fun w__287 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__287 with
+ | Some
+ (FCVT_S_WU, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_WU), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6035_
+ _s4974_) >>= fun w__290 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__290 with
+ | Some
+ (FCVT_L_S, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6035_
+ _s4974_) >>= fun w__291 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__291 with
+ | Some
+ (FCVT_L_S, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_L_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6052_
+ _s4974_) >>= fun w__294 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__294 with
+ | Some
+ (FCVT_LU_S, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6052_
+ _s4974_) >>= fun w__295 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__295 with
+ | Some
+ (FCVT_LU_S, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_LU_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6069_
+ _s4974_) >>= fun w__298 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__298 with
+ | Some
+ (FCVT_S_L, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6069_
+ _s4974_) >>= fun w__299 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__299 with
+ | Some
+ (FCVT_S_L, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_L), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6086_
+ _s4974_) >>= fun w__302 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__302 with
+ | Some
+ (FCVT_S_LU, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6086_
+ _s4974_) >>= fun w__303 : option ((f_un_rm_op_S * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__303 with
+ | Some
+ (FCVT_S_LU, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_S
+ (rs1, rm, rd, FCVT_S_LU), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6103_
+ _s4974_) >>= fun w__306 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__306 with
+ | Some
+ (FSGNJ_S, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6103_
+ _s4974_) >>= fun w__307 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__307 with
+ | Some
+ (FSGNJ_S, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJ_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6120_
+ _s4974_) >>= fun w__310 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__310 with
+ | Some
+ (FSGNJN_S, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6120_
+ _s4974_) >>= fun w__311 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__311 with
+ | Some
+ (FSGNJN_S, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJN_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6137_
+ _s4974_) >>= fun w__314 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__314 with
+ | Some
+ (FSGNJX_S, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6137_
+ _s4974_) >>= fun w__315 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__315 with
+ | Some
+ (FSGNJX_S, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_S
+ (rs2, rs1, rd, FSGNJX_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6154_
+ _s4974_) >>= fun w__318 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__318 with
+ | Some
+ (FMIN_S, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6154_
+ _s4974_) >>= fun w__319 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__319 with
+ | Some
+ (FMIN_S, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_S
+ (rs2, rs1, rd, FMIN_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6171_
+ _s4974_) >>= fun w__322 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__322 with
+ | Some
+ (FMAX_S, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6171_
+ _s4974_) >>= fun w__323 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__323 with
+ | Some
+ (FMAX_S, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_S
+ (rs2, rs1, rd, FMAX_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6188_
+ _s4974_) >>= fun w__326 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__326 with
+ | Some
+ (FEQ_S, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6188_
+ _s4974_) >>= fun w__327 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__327 with
+ | Some
+ (FEQ_S, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_S
+ (rs2, rs1, rd, FEQ_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6205_
+ _s4974_) >>= fun w__330 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__330 with
+ | Some
+ (FLT_S, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6205_
+ _s4974_) >>= fun w__331 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__331 with
+ | Some
+ (FLT_S, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_S
+ (rs2, rs1, rd, FLT_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6222_
+ _s4974_) >>= fun w__334 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__334 with
+ | Some
+ (FLE_S, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6222_
+ _s4974_) >>= fun w__335 : option ((f_bin_op_S * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__335 with
+ | Some
+ (FLE_S, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_S
+ (rs2, rs1, rd, FLE_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6239_
+ _s4974_) >>= fun w__338 : option ((f_un_op_S * mword 5 * mword 5 * string)) =>
+ (if
+ match w__338 with
+ | Some
+ (FMV_X_W, rd, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6239_
+ _s4974_) >>= fun w__339 : option ((f_un_op_S * mword 5 * mword 5 * string)) =>
+ (match w__339 with
+ | Some
+ (FMV_X_W, rd, rs1, s_) =>
+ returnm (Some
+ (F_UN_TYPE_S
+ (rs1, rd, FMV_X_W), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6252_
+ _s4974_) >>= fun w__342 : option ((f_un_op_S * mword 5 * mword 5 * string)) =>
+ (if
+ match w__342 with
+ | Some
+ (FMV_W_X, rd, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6252_
+ _s4974_) >>= fun w__343 : option ((f_un_op_S * mword 5 * mword 5 * string)) =>
+ (match w__343 with
+ | Some
+ (FMV_W_X, rd, rs1, s_) =>
+ returnm (Some
+ (F_UN_TYPE_S
+ (rs1, rd, FMV_W_X), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6265_
+ _s4974_) >>= fun w__346 : option ((f_un_op_S * mword 5 * mword 5 * string)) =>
+ (if
+ match w__346 with
+ | Some
+ (FCLASS_S, rd, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6265_
+ _s4974_) >>= fun w__347 : option ((f_un_op_S * mword 5 * mword 5 * string)) =>
+ (match w__347 with
+ | Some
+ (FCLASS_S, rd, rs1, s_) =>
+ returnm (Some
+ (F_UN_TYPE_S
+ (rs1, rd, FCLASS_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6278_
+ _s4974_) >>= fun w__350 : option ((mword 5 * mword 6 * string)) =>
+ (if
+ match w__350 with
+ | Some
+ (rd, imm, s_) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s6278_
+ _s4974_) >>= fun w__351 : option ((mword 5 * mword 6 * string)) =>
+ (match w__351 with
+ | Some
+ (rd, imm, s_) =>
+ returnm (Some
+ (C_FLWSP
+ (imm, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6290_
+ _s4974_) >>= fun w__354 : option ((mword 5 * mword 6 * string)) =>
+ (if
+ match w__354 with
+ | Some
+ (rd, uimm, s_) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s6290_
+ _s4974_) >>= fun w__355 : option ((mword 5 * mword 6 * string)) =>
+ (match w__355 with
+ | Some
+ (rd, uimm, s_) =>
+ returnm (Some
+ (C_FSWSP
+ (uimm, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6302_
+ _s4974_) >>= fun w__358 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (if
+ match w__358 with
+ | Some
+ (rdc, rsc, uimm, s_) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s6302_
+ _s4974_) >>= fun w__359 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (match w__359 with
+ | Some
+ (rdc, rsc, uimm, s_) =>
+ returnm (Some
+ (C_FLW
+ (uimm, rsc, rdc), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6318_
+ _s4974_) >>= fun w__362 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (if
+ match w__362 with
+ | Some
+ (rsc1, rsc2, uimm, s_) =>
+ Z.eqb
+ 64
+ 32
+ | _ =>
+ false
+ end
+ then
+ (_s6318_
+ _s4974_) >>= fun w__363 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (match w__363 with
+ | Some
+ (rsc1, rsc2, uimm, s_) =>
+ returnm (Some
+ (C_FSW
+ (uimm, rsc1, rsc2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6334_
+ _s4974_) >>= fun w__366 : option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__366 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6334_
+ _s4974_) >>= fun w__367 : option ((f_madd_op_D * mword 5 * mword 5 * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__367 with
+ | Some
+ (op, rd, rs1, rs2, rs3, rm, s_) =>
+ returnm (Some
+ (F_MADD_TYPE_D
+ (rs3, rs2, rs1, rm, rd, op), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6359_
+ _s4974_) >>= fun w__370 : option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__370 with
+ | Some
+ (op, rd, rs1, rs2, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6359_
+ _s4974_) >>= fun w__371 : option ((f_bin_rm_op_D * mword 5 * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__371 with
+ | Some
+ (op, rd, rs1, rs2, rm, s_) =>
+ returnm (Some
+ (F_BIN_RM_TYPE_D
+ (rs2, rs1, rm, rd, op), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6380_
+ _s4974_) >>= fun w__374 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__374 with
+ | Some
+ (FSQRT_D, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6380_
+ _s4974_) >>= fun w__375 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__375 with
+ | Some
+ (FSQRT_D, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FSQRT_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6397_
+ _s4974_) >>= fun w__378 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__378 with
+ | Some
+ (FCVT_W_D, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6397_
+ _s4974_) >>= fun w__379 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__379 with
+ | Some
+ (FCVT_W_D, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_W_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6414_
+ _s4974_) >>= fun w__382 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__382 with
+ | Some
+ (FCVT_WU_D, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6414_
+ _s4974_) >>= fun w__383 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__383 with
+ | Some
+ (FCVT_WU_D, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_WU_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6431_
+ _s4974_) >>= fun w__386 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__386 with
+ | Some
+ (FCVT_D_W, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6431_
+ _s4974_) >>= fun w__387 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__387 with
+ | Some
+ (FCVT_D_W, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_W), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6448_
+ _s4974_) >>= fun w__390 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__390 with
+ | Some
+ (FCVT_D_WU, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6448_
+ _s4974_) >>= fun w__391 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__391 with
+ | Some
+ (FCVT_D_WU, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_WU), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6465_
+ _s4974_) >>= fun w__394 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__394 with
+ | Some
+ (FCVT_L_D, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6465_
+ _s4974_) >>= fun w__395 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__395 with
+ | Some
+ (FCVT_L_D, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_L_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6482_
+ _s4974_) >>= fun w__398 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__398 with
+ | Some
+ (FCVT_LU_D, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6482_
+ _s4974_) >>= fun w__399 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__399 with
+ | Some
+ (FCVT_LU_D, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_LU_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6499_
+ _s4974_) >>= fun w__402 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__402 with
+ | Some
+ (FCVT_D_L, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6499_
+ _s4974_) >>= fun w__403 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__403 with
+ | Some
+ (FCVT_D_L, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_L), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6516_
+ _s4974_) >>= fun w__406 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__406 with
+ | Some
+ (FCVT_D_LU, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6516_
+ _s4974_) >>= fun w__407 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__407 with
+ | Some
+ (FCVT_D_LU, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_LU), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6533_
+ _s4974_) >>= fun w__410 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__410 with
+ | Some
+ (FCVT_S_D, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6533_
+ _s4974_) >>= fun w__411 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__411 with
+ | Some
+ (FCVT_S_D, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_S_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6550_
+ _s4974_) >>= fun w__414 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (if
+ match w__414 with
+ | Some
+ (FCVT_D_S, rd, rs1, rm, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6550_
+ _s4974_) >>= fun w__415 : option ((f_un_rm_op_D * mword 5 * mword 5 * rounding_mode * string)) =>
+ (match w__415 with
+ | Some
+ (FCVT_D_S, rd, rs1, rm, s_) =>
+ returnm (Some
+ (F_UN_RM_TYPE_D
+ (rs1, rm, rd, FCVT_D_S), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6567_
+ _s4974_) >>= fun w__418 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__418 with
+ | Some
+ (FSGNJ_D, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6567_
+ _s4974_) >>= fun w__419 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__419 with
+ | Some
+ (FSGNJ_D, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJ_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6584_
+ _s4974_) >>= fun w__422 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__422 with
+ | Some
+ (FSGNJN_D, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6584_
+ _s4974_) >>= fun w__423 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__423 with
+ | Some
+ (FSGNJN_D, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJN_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6601_
+ _s4974_) >>= fun w__426 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__426 with
+ | Some
+ (FSGNJX_D, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6601_
+ _s4974_) >>= fun w__427 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__427 with
+ | Some
+ (FSGNJX_D, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_D
+ (rs2, rs1, rd, FSGNJX_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6618_
+ _s4974_) >>= fun w__430 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__430 with
+ | Some
+ (FMIN_D, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6618_
+ _s4974_) >>= fun w__431 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__431 with
+ | Some
+ (FMIN_D, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_D
+ (rs2, rs1, rd, FMIN_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6635_
+ _s4974_) >>= fun w__434 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__434 with
+ | Some
+ (FMAX_D, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6635_
+ _s4974_) >>= fun w__435 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__435 with
+ | Some
+ (FMAX_D, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_D
+ (rs2, rs1, rd, FMAX_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6652_
+ _s4974_) >>= fun w__438 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__438 with
+ | Some
+ (FEQ_D, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6652_
+ _s4974_) >>= fun w__439 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__439 with
+ | Some
+ (FEQ_D, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_D
+ (rs2, rs1, rd, FEQ_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6669_
+ _s4974_) >>= fun w__442 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__442 with
+ | Some
+ (FLT_D, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6669_
+ _s4974_) >>= fun w__443 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__443 with
+ | Some
+ (FLT_D, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_D
+ (rs2, rs1, rd, FLT_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6686_
+ _s4974_) >>= fun w__446 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (if
+ match w__446 with
+ | Some
+ (FLE_D, rd, rs1, rs2, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6686_
+ _s4974_) >>= fun w__447 : option ((f_bin_op_D * mword 5 * mword 5 * mword 5 * string)) =>
+ (match w__447 with
+ | Some
+ (FLE_D, rd, rs1, rs2, s_) =>
+ returnm (Some
+ (F_BIN_TYPE_D
+ (rs2, rs1, rd, FLE_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6703_
+ _s4974_) >>= fun w__450 : option ((f_un_op_D * mword 5 * mword 5 * string)) =>
+ (if
+ match w__450 with
+ | Some
+ (FMV_X_D, rd, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6703_
+ _s4974_) >>= fun w__451 : option ((f_un_op_D * mword 5 * mword 5 * string)) =>
+ (match w__451 with
+ | Some
+ (FMV_X_D, rd, rs1, s_) =>
+ returnm (Some
+ (F_UN_TYPE_D
+ (rs1, rd, FMV_X_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6716_
+ _s4974_) >>= fun w__454 : option ((f_un_op_D * mword 5 * mword 5 * string)) =>
+ (if
+ match w__454 with
+ | Some
+ (FMV_D_X, rd, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6716_
+ _s4974_) >>= fun w__455 : option ((f_un_op_D * mword 5 * mword 5 * string)) =>
+ (match w__455 with
+ | Some
+ (FMV_D_X, rd, rs1, s_) =>
+ returnm (Some
+ (F_UN_TYPE_D
+ (rs1, rd, FMV_D_X), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6729_
+ _s4974_) >>= fun w__458 : option ((f_un_op_D * mword 5 * mword 5 * string)) =>
+ (if
+ match w__458 with
+ | Some
+ (FCLASS_D, rd, rs1, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6729_
+ _s4974_) >>= fun w__459 : option ((f_un_op_D * mword 5 * mword 5 * string)) =>
+ (match w__459 with
+ | Some
+ (FCLASS_D, rd, rs1, s_) =>
+ returnm (Some
+ (F_UN_TYPE_D
+ (rs1, rd, FCLASS_D), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6742_
+ _s4974_) >>= fun w__462 : option ((mword 5 * mword 6 * string)) =>
+ (if
+ match w__462 with
+ | Some
+ (rd, uimm, s_) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s6742_
+ _s4974_) >>= fun w__463 : option ((mword 5 * mword 6 * string)) =>
+ (match w__463 with
+ | Some
+ (rd, uimm, s_) =>
+ returnm (Some
+ (C_FLDSP
+ (uimm, rd), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6754_
+ _s4974_) >>= fun w__466 : option ((mword 5 * mword 6 * string)) =>
+ (if
+ match w__466 with
+ | Some
+ (rs2, uimm, s_) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s6754_
+ _s4974_) >>= fun w__467 : option ((mword 5 * mword 6 * string)) =>
+ (match w__467 with
+ | Some
+ (rs2, uimm, s_) =>
+ returnm (Some
+ (C_FSDSP
+ (uimm, rs2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6766_
+ _s4974_) >>= fun w__470 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (if
+ match w__470 with
+ | Some
+ (rdc, rsc, uimm, s_) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s6766_
+ _s4974_) >>= fun w__471 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (match w__471 with
+ | Some
+ (rdc, rsc, uimm, s_) =>
+ returnm (Some
+ (C_FLD
+ (uimm, rsc, rdc), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6782_
+ _s4974_) >>= fun w__474 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (if
+ match w__474 with
+ | Some
+ (rsc1, rsc2, uimm, s_) =>
+ orb
+ (Z.eqb
+ 64
+ 32)
+ (Z.eqb
+ 64
+ 64)
+ | _ =>
+ false
+ end
+ then
+ (_s6782_
+ _s4974_) >>= fun w__475 : option ((mword 3 * mword 3 * mword 5 * string)) =>
+ (match w__475 with
+ | Some
+ (rsc1, rsc2, uimm, s_) =>
+ returnm (Some
+ (C_FSD
+ (uimm, rsc1, rsc2), build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6798_
+ _s4974_) >>= fun w__478 : option ((mword 32 * string)) =>
+ (if
+ match w__478 with
+ | Some
+ (s, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6798_
+ _s4974_) >>= fun w__479 : option ((mword 32 * string)) =>
+ (match w__479 with
+ | Some
+ (s, s_) =>
+ returnm (Some
+ (ILLEGAL
+ s, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ (_s6806_
+ _s4974_) >>= fun w__482 : option ((mword 16 * string)) =>
+ (if
+ match w__482 with
+ | Some
+ (s, s_) =>
+ true
+ | _ =>
+ false
+ end
+ then
+ (_s6806_
+ _s4974_) >>= fun w__483 : option ((mword 16 * string)) =>
+ (match w__483 with
+ | Some
+ (s, s_) =>
+ returnm (Some
+ (C_ILLEGAL
+ s, build_ex
+ (projT1
+ (sub_nat
+ (projT1
+ (string_length
+ arg_))
+ (projT1
+ (string_length
+ s_))))))
+ | _ =>
+ exit tt
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ end)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)})))
+ else
+ returnm None)
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
+ 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >=
+ : M (option ((ast * {n : Z & ArithFact (n >=?
0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))))
- : M (option ((ast * {n : Z & ArithFact (n >= 0)}))).
-
-Definition print_insn (insn : ast) : M (string) := (assembly_forwards insn) : M (string).
-
-Definition decode (bv : mword 32) : M (ast) := (encdec_backwards bv) : M (ast).
-
-Definition decodeCompressed (bv : mword 16) : ast := encdec_compressed_backwards bv.
-
-Definition GPRstr : vec string 32 :=
-vec_of_list_len ["x31";"x30";"x29";"x28";"x27";"x26";"x25";"x24";"x23";"x22";"x21";"x20";"x19";"x18";"x17";"x16";"x15";"x14";"x13";"x12";"x11";
- "x10";"x9";"x8";"x7";"x6";"x5";"x4";"x3";"x2";"x1";"x0"].
-Hint Unfold GPRstr : sail.
-Definition CIA_fp := RFull ("CIA").
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))))
+ : M (option ((ast * {n : Z & ArithFact (n >=? 0)}))).
+
+Definition print_insn (insn : ast) : M (string) := (assembly_forwards insn) : M (string).
+
+Definition decode (bv : mword 32) : M (ast) := (encdec_backwards bv) : M (ast).
+
+Definition decodeCompressed (bv : mword 16) : M (ast) :=
+ (encdec_compressed_backwards bv) : M (ast).
+
+Definition GPRstrs : vec string 32 :=
+vec_of_list_len ["x31";"x30";"x29";"x28";"x27";"x26";"x25";"x24";"x23";"x22";"x21";"x20";"x19";"x18";
+ "x17";"x16";"x15";"x14";"x13";"x12";"x11";"x10";"x9";"x8";"x7";"x6";"x5";"x4";"x3";
+ "x2";"x1";"x0"].
+Hint Unfold GPRstrs : sail.
+Definition GPRstr (i : mword 5) : string := vec_access_dec GPRstrs (projT1 (uint i)).
+
+Definition CIA_fp := RFull "CIA".
Hint Unfold CIA_fp : sail.
-Definition NIA_fp := RFull ("NIA").
+Definition NIA_fp := RFull "NIA".
Hint Unfold NIA_fp : sail.
-Definition initial_analysis (instr : ast)
+Definition initial_analysis (instr : ast)
: M ((list regfp * list regfp * list regfp * list niafp * diafp * instruction_kind)) :=
-
let iR := [] : regfps in
let oR := [] : regfps in
let aR := [] : regfps in
- let ik := (IK_simple (tt)) : instruction_kind in
- let Nias := [NIAFP_successor (tt)] : niafps in
- let Dia := (DIAFP_none (tt)) : diafp in
+ let ik := (IK_simple tt) : instruction_kind in
+ let Nias := [NIAFP_successor tt] : niafps in
+ let Dia := (DIAFP_none tt) : diafp in
(match instr with
- | EBREAK (tt) => returnm (Nias, aR, iR, ik, oR)
- | UTYPE ((imm, rd, op)) =>
+ | EBREAK tt => returnm (Nias, aR, iR, ik, oR)
+ | UTYPE (imm, rd, op) =>
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
returnm (Nias, aR, iR, ik, oR)
- | RISCV_JAL ((imm, rd)) =>
+ | RISCV_JAL (imm, rd) =>
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
let offset : bits 64 := EXTS 64 imm in
((read_reg PC_ref) : M (mword 64)) >>= fun w__0 : mword 64 =>
let Nias : list niafp := [NIAFP_concrete_address (add_vec w__0 offset)] in
- let ik : instruction_kind := IK_branch (tt) in
+ let ik : instruction_kind := IK_branch tt in
returnm (Nias, aR, iR, ik, oR)
- | RISCV_JALR ((imm, rs, rd)) =>
+ | RISCV_JALR (imm, rs, rd) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs)))) :: iR in
+ if eq_vec rs ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
let offset : bits 64 := EXTS 64 imm in
- let Nias : list niafp := [NIAFP_indirect_address (tt)] in
- let ik : instruction_kind := IK_branch (tt) in
+ let Nias : list niafp := [NIAFP_indirect_address tt] in
+ let ik : instruction_kind := IK_branch tt in
returnm (Nias, aR, iR, ik, oR)
- | BTYPE ((imm, rs2, rs1, op)) =>
+ | BTYPE (imm, rs2, rs1, op) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs2)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs2)))) :: iR in
+ if eq_vec rs2 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs2)) :: iR in
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: iR in
- let ik := (IK_branch (tt)) : instruction_kind in
+ if eq_vec rs1 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs1)) :: iR in
+ let ik := (IK_branch tt) : instruction_kind in
let offset : bits 64 := EXTS 64 imm in
((read_reg PC_ref) : M (mword 64)) >>= fun w__1 : mword 64 =>
- let Nias : list niafp := [NIAFP_concrete_address (add_vec w__1 offset);NIAFP_successor (tt)] in
+ let Nias : list niafp := [NIAFP_concrete_address (add_vec w__1 offset); NIAFP_successor tt] in
returnm (Nias, aR, iR, ik, oR)
- | ITYPE ((imm, rs, rd, op)) =>
+ | ITYPE (imm, rs, rd, op) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs)))) :: iR in
+ if eq_vec rs ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
returnm (Nias, aR, iR, ik, oR)
- | SHIFTIOP ((imm, rs, rd, op)) =>
+ | SHIFTIOP (imm, rs, rd, op) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs)))) :: iR in
+ if eq_vec rs ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
returnm (Nias, aR, iR, ik, oR)
- | RTYPE ((rs2, rs1, rd, op)) =>
+ | RTYPE (rs2, rs1, rd, op) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs2)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs2)))) :: iR in
+ if eq_vec rs2 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs2)) :: iR in
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: iR in
+ if eq_vec rs1 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs1)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
returnm (Nias, aR, iR, ik, oR)
- | CSR ((csr, rs1, rd, is_imm, op)) =>
+ | CSR (csr, rs1, rd, is_imm, op) =>
let isWrite : bool :=
match op with
| CSRRW => true
| _ =>
- if sumbool_of_bool (is_imm) then projT1 (neq_int (projT1 (uint rs1)) 0)
+ if sumbool_of_bool is_imm then projT1 (neq_int (projT1 (uint rs1)) 0)
else projT1 (neq_int (projT1 (uint rs1)) 0)
end in
let iR : list regfp := (RFull (csr_name csr)) :: iR in
let iR : list regfp :=
- if sumbool_of_bool ((negb is_imm)) then
- (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: iR
+ if sumbool_of_bool (negb is_imm) then (RFull (GPRstr rs1)) :: iR
else iR in
- let oR : list regfp := if sumbool_of_bool (isWrite) then (RFull (csr_name csr)) :: oR else oR in
- let oR : list regfp := (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ let oR : list regfp := if sumbool_of_bool isWrite then (RFull (csr_name csr)) :: oR else oR in
+ let oR : list regfp := (RFull (GPRstr rd)) :: oR in
returnm (Nias, aR, iR, ik, oR)
- | LOAD ((imm, rs, rd, unsign, width, aq, rl)) =>
+ | LOAD (imm, rs, rd, unsign, width, aq, rl) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs)))) :: iR in
+ if eq_vec rs ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
let aR := iR : list regfp in
(match (aq, rl) with
- | (false, false) => returnm ((IK_mem_read (Read_plain)) : instruction_kind )
- | (true, false) => returnm ((IK_mem_read (Read_RISCV_acquire)) : instruction_kind )
- | (true, true) => returnm ((IK_mem_read (Read_RISCV_strong_acquire)) : instruction_kind )
+ | (false, false) => returnm (IK_mem_read Read_plain)
+ | (true, false) => returnm (IK_mem_read Read_RISCV_acquire)
+ | (true, true) => returnm (IK_mem_read Read_RISCV_strong_acquire)
| _ =>
(internal_error "LOAD type not implemented in initial_analysis") : M (instruction_kind)
end) >>= fun w__3 : instruction_kind =>
let ik : instruction_kind := w__3 in
returnm (Nias, aR, iR, ik, oR)
- | STORE ((imm, rs2, rs1, width, aq, rl)) =>
+ | STORE (imm, rs2, rs1, width, aq, rl) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs2)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs2)))) :: iR in
+ if eq_vec rs2 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs2)) :: iR in
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: iR in
+ if eq_vec rs1 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs1)) :: iR in
let aR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then aR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: aR in
+ if eq_vec rs1 ('b"00000" : mword 5) then aR
+ else (RFull (GPRstr rs1)) :: aR in
(match (aq, rl) with
- | (false, false) => returnm ((IK_mem_write (Write_plain)) : instruction_kind )
- | (false, true) => returnm ((IK_mem_write (Write_RISCV_release)) : instruction_kind )
- | (true, true) => returnm ((IK_mem_write (Write_RISCV_strong_release)) : instruction_kind )
+ | (false, false) => returnm (IK_mem_write Write_plain)
+ | (false, true) => returnm (IK_mem_write Write_RISCV_release)
+ | (true, true) => returnm (IK_mem_write Write_RISCV_strong_release)
| _ =>
(internal_error "STORE type not implemented in initial_analysis") : M (instruction_kind)
end) >>= fun w__5 : instruction_kind =>
let ik : instruction_kind := w__5 in
returnm (Nias, aR, iR, ik, oR)
- | ADDIW ((imm, rs, rd)) =>
+ | ADDIW (imm, rs, rd) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs)))) :: iR in
+ if eq_vec rs ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
returnm (Nias, aR, iR, ik, oR)
- | SHIFTW ((imm, rs, rd, op)) =>
+ | SHIFTW (imm, rs, rd, op) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs)))) :: iR in
+ if eq_vec rs ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
returnm (Nias, aR, iR, ik, oR)
- | RTYPEW ((rs2, rs1, rd, op)) =>
+ | RTYPEW (rs2, rs1, rd, op) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs2)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs2)))) :: iR in
+ if eq_vec rs2 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs2)) :: iR in
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: iR in
+ if eq_vec rs1 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs1)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
returnm (Nias, aR, iR, ik, oR)
- | FENCE ((pred, succ)) =>
+ | FENCE (pred, succ) =>
(match (pred, succ) with
- | (v__838, v__839) =>
- (if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))))
- then
- returnm ((IK_barrier
- (Barrier_RISCV_rw_rw
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B1;B1] : mword (1 - 0 + 1))))) then
- returnm ((IK_barrier
- (Barrier_RISCV_r_rw
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1))))) then
- returnm ((IK_barrier
- (Barrier_RISCV_r_r
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))) then
- returnm ((IK_barrier
- (Barrier_RISCV_rw_w
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))) then
- returnm ((IK_barrier
- (Barrier_RISCV_w_w
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B1;B1] : mword (1 - 0 + 1))))) then
- returnm ((IK_barrier
- (Barrier_RISCV_w_rw
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1))))) then
- returnm ((IK_barrier
- (Barrier_RISCV_rw_r
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1))))) then
- returnm ((IK_barrier
- (Barrier_RISCV_r_w
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B0;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B1;B0] : mword (1 - 0 + 1))))) then
- returnm ((IK_barrier
- (Barrier_RISCV_w_r
- (tt)))
- : instruction_kind )
- else if ((andb
- (eq_vec (subrange_vec_dec v__838 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__839 1 0)
- (vec_of_bits [B0;B0] : mword (1 - 0 + 1))))) then
- returnm ((IK_simple
- (tt))
- : instruction_kind )
+ | (v__1398, v__1399) =>
+ (if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"11" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"11" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_rw_rw tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"10" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"11" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_r_rw tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"10" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"10" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_r_r tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"11" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_rw_w tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"01" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_w_w tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"01" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"11" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_w_rw tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"11" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"10" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_rw_r tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"10" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"01" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_r_w tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"01" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"10" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_w_r tt))
+ else if andb (eq_vec (subrange_vec_dec v__1398 1 0) ('b"00" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1399 1 0) ('b"00" : mword (1 - 0 + 1))) then
+ returnm (IK_simple tt)
else
(internal_error "barrier type not implemented in initial_analysis")
: M (instruction_kind))
@@ -42910,17 +67525,12 @@ Definition initial_analysis (instr : ast)
end) >>= fun w__17 : instruction_kind =>
let ik : instruction_kind := w__17 in
returnm (Nias, aR, iR, ik, oR)
- | FENCE_TSO ((pred, succ)) =>
+ | FENCE_TSO (pred, succ) =>
(match (pred, succ) with
- | (v__878, v__879) =>
- (if ((andb
- (eq_vec (subrange_vec_dec v__878 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))
- (eq_vec (subrange_vec_dec v__879 1 0) (vec_of_bits [B1;B1] : mword (1 - 0 + 1)))))
- then
- returnm ((IK_barrier
- (Barrier_RISCV_tso
- (tt)))
- : instruction_kind )
+ | (v__1438, v__1439) =>
+ (if andb (eq_vec (subrange_vec_dec v__1438 1 0) ('b"11" : mword (1 - 0 + 1)))
+ (eq_vec (subrange_vec_dec v__1439 1 0) ('b"11" : mword (1 - 0 + 1))) then
+ returnm (IK_barrier (Barrier_RISCV_tso tt))
else
(internal_error "barrier type not implemented in initial_analysis")
: M (instruction_kind))
@@ -42928,74 +67538,69 @@ Definition initial_analysis (instr : ast)
end) >>= fun w__20 : instruction_kind =>
let ik : instruction_kind := w__20 in
returnm (Nias, aR, iR, ik, oR)
- | FENCEI (tt) =>
- let ik : instruction_kind := IK_simple (tt) in
+ | FENCEI tt =>
+ let ik : instruction_kind := IK_simple tt in
returnm (Nias, aR, iR, ik, oR)
- | LOADRES ((aq, rl, rs1, width, rd)) =>
+ | LOADRES (aq, rl, rs1, width, rd) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: iR in
+ if eq_vec rs1 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs1)) :: iR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
let aR := iR : list regfp in
(match (aq, rl) with
- | (false, false) => returnm ((IK_mem_read (Read_RISCV_reserved)) : instruction_kind )
- | (true, false) =>
- returnm ((IK_mem_read (Read_RISCV_reserved_acquire)) : instruction_kind )
- | (true, true) =>
- returnm ((IK_mem_read (Read_RISCV_reserved_strong_acquire)) : instruction_kind )
+ | (false, false) => returnm (IK_mem_read Read_RISCV_reserved)
+ | (true, false) => returnm (IK_mem_read Read_RISCV_reserved_acquire)
+ | (true, true) => returnm (IK_mem_read Read_RISCV_reserved_strong_acquire)
| (false, true) =>
(internal_error "LOADRES type not implemented in initial_analysis")
: M (instruction_kind)
end) >>= fun w__22 : instruction_kind =>
let ik : instruction_kind := w__22 in
returnm (Nias, aR, iR, ik, oR)
- | STORECON ((aq, rl, rs2, rs1, width, rd)) =>
+ | STORECON (aq, rl, rs2, rs1, width, rd) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs2)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs2)))) :: iR in
+ if eq_vec rs2 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs2)) :: iR in
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: iR in
+ if eq_vec rs1 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs1)) :: iR in
let aR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then aR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: aR in
+ if eq_vec rs1 ('b"00000" : mword 5) then aR
+ else (RFull (GPRstr rs1)) :: aR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
(match (aq, rl) with
- | (false, false) => returnm ((IK_mem_write (Write_RISCV_conditional)) : instruction_kind )
- | (false, true) =>
- returnm ((IK_mem_write (Write_RISCV_conditional_release)) : instruction_kind )
- | (true, true) =>
- returnm ((IK_mem_write (Write_RISCV_conditional_strong_release)) : instruction_kind )
+ | (false, false) => returnm (IK_mem_write Write_RISCV_conditional)
+ | (false, true) => returnm (IK_mem_write Write_RISCV_conditional_release)
+ | (true, true) => returnm (IK_mem_write Write_RISCV_conditional_strong_release)
| (true, false) =>
(internal_error "STORECON type not implemented in initial_analysis")
: M (instruction_kind)
end) >>= fun w__24 : instruction_kind =>
let ik : instruction_kind := w__24 in
returnm (Nias, aR, iR, ik, oR)
- | AMO ((op, aq, rl, rs2, rs1, width, rd)) =>
+ | AMO (op, aq, rl, rs2, rs1, width, rd) =>
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs2)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs2)))) :: iR in
+ if eq_vec rs2 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs2)) :: iR in
let iR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then iR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: iR in
+ if eq_vec rs1 ('b"00000" : mword 5) then iR
+ else (RFull (GPRstr rs1)) :: iR in
let aR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rs1)) 0)) then aR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rs1)))) :: aR in
+ if eq_vec rs1 ('b"00000" : mword 5) then aR
+ else (RFull (GPRstr rs1)) :: aR in
let oR : list regfp :=
- if sumbool_of_bool ((Z.eqb (projT1 (regidx_to_regno rd)) 0)) then oR
- else (RFull (vec_access_dec GPRstr (projT1 (regidx_to_regno rd)))) :: oR in
+ if eq_vec rd ('b"00000" : mword 5) then oR
+ else (RFull (GPRstr rd)) :: oR in
let ik : instruction_kind :=
match (aq, rl) with
- | (false, false) => IK_mem_rmw ((Read_RISCV_reserved, Write_RISCV_conditional))
- | (false, true) => IK_mem_rmw ((Read_RISCV_reserved, Write_RISCV_conditional_release))
- | (true, false) => IK_mem_rmw ((Read_RISCV_reserved_acquire, Write_RISCV_conditional))
- | (true, true) =>
- IK_mem_rmw ((Read_RISCV_reserved_acquire, Write_RISCV_conditional_release))
+ | (false, false) => IK_mem_rmw (Read_RISCV_reserved, Write_RISCV_conditional)
+ | (false, true) => IK_mem_rmw (Read_RISCV_reserved, Write_RISCV_conditional_release)
+ | (true, false) => IK_mem_rmw (Read_RISCV_reserved_acquire, Write_RISCV_conditional)
+ | (true, true) => IK_mem_rmw (Read_RISCV_reserved_acquire, Write_RISCV_conditional_release)
end in
returnm (Nias, aR, iR, ik, oR)
| _ => returnm (Nias, aR, iR, ik, oR)
@@ -43004,4 +67609,4 @@ Definition initial_analysis (instr : ast)
returnm (iR, oR, aR, Nias, Dia, ik).
-End Content.
+
diff --git a/prover_snapshots/coq/RV64/riscv_extras.v b/prover_snapshots/coq/RV64/riscv_extras.v
index 84f6761..02d5609 100644
--- a/prover_snapshots/coq/RV64/riscv_extras.v
+++ b/prover_snapshots/coq/RV64/riscv_extras.v
@@ -1,13 +1,8 @@
-Require Import Sail2_instr_kinds.
-Require Import Sail2_values.
-Require Import Sail2_operators_mwords.
-Require Import Sail2_prompt_monad.
-Require Import Sail2_prompt.
+Require Import Sail.Base.
Require Import String.
Require Import List.
Import List.ListNotations.
-
-Axiom real : Type.
+Open Scope Z.
Definition MEM_fence_rw_rw {rv e} (_:unit) : monad rv unit e := barrier (Barrier_RISCV_rw_rw tt).
Definition MEM_fence_r_rw {rv e} (_:unit) : monad rv unit e := barrier (Barrier_RISCV_r_rw tt).
@@ -45,12 +40,12 @@ val MEMr_reserved_acquire : forall 'rv 'a 'b 'e. Size 'a, Size 'b => inte
val MEMr_reserved_strong_acquire : forall 'rv 'a 'b 'e. Size 'a, Size 'b => integer -> integer -> bitvector 'a -> bitvector 'a -> monad 'rv (bitvector 'b) 'e
*)
-Definition MEMr {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >= 0)} : monad rv (mword (8 * size)) e := read_mem Read_plain addrsize addr size.
-Definition MEMr_acquire {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >= 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_acquire addrsize addr size.
-Definition MEMr_strong_acquire {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >= 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_strong_acquire addrsize addr size.
-Definition MEMr_reserved {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >= 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_reserved addrsize addr size.
-Definition MEMr_reserved_acquire {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >= 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_reserved_acquire addrsize addr size.
-Definition MEMr_reserved_strong_acquire {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >= 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_reserved_strong_acquire addrsize addr size.
+Definition MEMr {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >=? 0)} : monad rv (mword (8 * size)) e := read_mem Read_plain addrsize addr size.
+Definition MEMr_acquire {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >=? 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_acquire addrsize addr size.
+Definition MEMr_strong_acquire {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >=? 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_strong_acquire addrsize addr size.
+Definition MEMr_reserved {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >=? 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_reserved addrsize addr size.
+Definition MEMr_reserved_acquire {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >=? 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_reserved_acquire addrsize addr size.
+Definition MEMr_reserved_strong_acquire {rv e} addrsize size (hexRAM addr : mword addrsize) `{ArithFact (size >=? 0)} : monad rv (mword (8 * size)) e := read_mem Read_RISCV_reserved_strong_acquire addrsize addr size.
(*
val MEMw : forall 'rv 'a 'b 'e. Size 'a, Size 'b => integer -> integer -> bitvector 'a -> bitvector 'a -> bitvector 'b -> monad 'rv bool 'e
@@ -87,14 +82,14 @@ Definition undefined_string {rv e} (_:unit) : monad rv string e := returnm ""%st
Definition undefined_unit {rv e} (_:unit) : monad rv unit e := returnm tt.
Definition undefined_int {rv e} (_:unit) : monad rv Z e := returnm (0:ii).
(*val undefined_vector : forall 'rv 'a 'e. integer -> 'a -> monad 'rv (list 'a) 'e*)
-Definition undefined_vector {rv a e} len (u : a) `{ArithFact (len >= 0)} : monad rv (vec a len) e := returnm (vec_init u len).
+Definition undefined_vector {rv a e} len (u : a) `{ArithFact (len >=? 0)} : monad rv (vec a len) e := returnm (vec_init u len).
(*val undefined_bitvector : forall 'rv 'a 'e. Bitvector 'a => integer -> monad 'rv 'a 'e*)
-Definition undefined_bitvector {rv e} len `{ArithFact (len >= 0)} : monad rv (mword len) e := returnm (mword_of_int 0).
+Definition undefined_bitvector {rv e} len `{ArithFact (len >=? 0)} : monad rv (mword len) e := returnm (mword_of_int 0).
(*val undefined_bits : forall 'rv 'a 'e. Bitvector 'a => integer -> monad 'rv 'a 'e*)
Definition undefined_bits {rv e} := @undefined_bitvector rv e.
Definition undefined_bit {rv e} (_:unit) : monad rv bitU e := returnm BU.
(*Definition undefined_real {rv e} (_:unit) : monad rv real e := returnm (realFromFrac 0 1).*)
-Definition undefined_range {rv e} i j `{ArithFact (i <= j)} : monad rv {z : Z & ArithFact (i <= z /\ z <= j)} e := returnm (build_ex i).
+Definition undefined_range {rv e} i j `{ArithFact (i <=? j)} : monad rv {z : Z & ArithFact (i <=? z <=? j)} e := returnm (build_ex i).
Definition undefined_atom {rv e} i : monad rv Z e := returnm i.
Definition undefined_nat {rv e} (_:unit) : monad rv Z e := returnm (0:ii).
@@ -119,10 +114,12 @@ Definition eq_bit (x : bitU) (y : bitU) : bool :=
end.
Require Import Zeuclid.
-Definition euclid_modulo (m n : Z) `{ArithFact (n > 0)} : {z : Z & ArithFact (0 <= z <= n-1)}.
+Definition euclid_modulo (m n : Z) `{ArithFact (n >? 0)} : {z : Z & ArithFact (0 <=? z <=? n-1)}.
apply existT with (x := ZEuclid.modulo m n).
constructor.
destruct H.
+unbool_comparisons.
+unbool_comparisons_goal.
assert (Z.abs n = n). { rewrite Z.abs_eq; auto with zarith. }
rewrite <- H at 3.
lapply (ZEuclid.mod_always_pos m n); omega.
@@ -143,6 +140,7 @@ Definition string_of_int z := DecimalString.NilZero.string_of_int (Z.to_int z).
Axiom sys_enable_writable_misa : unit -> bool.
Axiom sys_enable_rvc : unit -> bool.
+Axiom sys_enable_fdext : unit -> bool.
(* The constraint solver can do this itself, but a Coq bug puts
anonymous_subproof into the term instead of an actual subproof. *)
diff --git a/prover_snapshots/coq/RV64/riscv_types.v b/prover_snapshots/coq/RV64/riscv_types.v
index 98f6666..a68c649 100644
--- a/prover_snapshots/coq/RV64/riscv_types.v
+++ b/prover_snapshots/coq/RV64/riscv_types.v
@@ -1,19 +1,17 @@
(*Generated by Sail from riscv.*)
-Require Import Sail2_instr_kinds.
-Require Import Sail2_values.
-Require Import Sail2_string.
-Require Import Sail2_real.
-Require Import Sail2_operators_mwords.
-Require Import Sail2_prompt_monad.
-Require Import Sail2_prompt.
-Require Import Sail2_state.
+Require Import Sail.Base.
+Require Import Sail.Real.
Import ListNotations.
+Open Scope string.
+Open Scope bool.
+Open Scope Z.
+
Definition bits (n : Z) : Type := mword n.
Inductive regfp :=
| RFull : string -> regfp
- | RSlice : (string * {n : Z & ArithFact (n >= 0)} * {n : Z & ArithFact (n >= 0)}) -> regfp
- | RSliceBit : (string * {n : Z & ArithFact (n >= 0)}) -> regfp
+ | RSlice : (string * {n : Z & ArithFact (n >=? 0)} * {n : Z & ArithFact (n >=? 0)}) -> regfp
+ | RSliceBit : (string * {n : Z & ArithFact (n >=? 0)}) -> regfp
| RField : (string * string) -> regfp.
Arguments regfp : clear implicits.
@@ -68,11 +66,27 @@ Hint Unfold xlen_bytes : sail.
Definition xlenbits : Type := bits 64.
+Definition flen : Z := 64.
+Hint Unfold flen : sail.
+
+Definition flen_bytes : Z := 8.
+Hint Unfold flen_bytes : sail.
+
+Definition flenbits : Type := bits 64.
+
Definition mem_meta : Type := unit.
Definition max_mem_access : Z := 16.
Hint Unfold max_mem_access : sail.
+Definition exc_code : Type := bits 8.
+
+Definition ext_ptw : Type := unit.
+
+Definition ext_ptw_error : Type := unit.
+
+Definition ext_exc_type : Type := unit.
+
Definition half : Type := bits 16.
Definition word : Type := bits 32.
@@ -83,7 +97,7 @@ Definition cregidx : Type := bits 3.
Definition csreg : Type := bits 12.
-Definition regno (n : Z)`{ArithFact (0 <= n /\ n < 32)} : Type := Z.
+Definition regno (n : Z)`{ArithFact ((0 <=? n) && (n <? 32))} : Type := Z.
Definition opcode : Type := bits 7.
@@ -109,6 +123,38 @@ Instance Decidable_eq_Privilege :
forall (x y : Privilege), Decidable (x = y) :=
Decidable_eq_from_dec Privilege_eq_dec.
+Inductive AccessType {a : Type} :=
+ | Read : a -> AccessType
+ | Write : a -> AccessType
+ | ReadWrite : a -> AccessType
+ | Execute : unit -> AccessType.
+Arguments AccessType : clear implicits.
+Instance Decidable_eq_AccessType {a : Type} `{forall x y : a, Decidable (x = y)}:
+forall (x y : AccessType a), Decidable (x = y).
+refine (Decidable_eq_from_dec (fun x y => _)).
+decide equality; refine (generic_dec _ _).
+Defined.
+
+Inductive ExceptionType :=
+ | E_Fetch_Addr_Align : unit -> ExceptionType
+ | E_Fetch_Access_Fault : unit -> ExceptionType
+ | E_Illegal_Instr : unit -> ExceptionType
+ | E_Breakpoint : unit -> ExceptionType
+ | E_Load_Addr_Align : unit -> ExceptionType
+ | E_Load_Access_Fault : unit -> ExceptionType
+ | E_SAMO_Addr_Align : unit -> ExceptionType
+ | E_SAMO_Access_Fault : unit -> ExceptionType
+ | E_U_EnvCall : unit -> ExceptionType
+ | E_S_EnvCall : unit -> ExceptionType
+ | E_Reserved_10 : unit -> ExceptionType
+ | E_M_EnvCall : unit -> ExceptionType
+ | E_Fetch_Page_Fault : unit -> ExceptionType
+ | E_Load_Page_Fault : unit -> ExceptionType
+ | E_Reserved_14 : unit -> ExceptionType
+ | E_SAMO_Page_Fault : unit -> ExceptionType
+ | E_Extension : ext_exc_type -> ExceptionType.
+Arguments ExceptionType : clear implicits.
+
Inductive amoop := AMOSWAP | AMOADD | AMOXOR | AMOAND | AMOOR | AMOMIN | AMOMAX | AMOMINU | AMOMAXU.
Scheme Equality for amoop.
Instance Decidable_eq_amoop :
@@ -127,6 +173,86 @@ Instance Decidable_eq_csrop :
forall (x y : csrop), Decidable (x = y) :=
Decidable_eq_from_dec csrop_eq_dec.
+Inductive f_bin_op_D := FSGNJ_D | FSGNJN_D | FSGNJX_D | FMIN_D | FMAX_D | FEQ_D | FLT_D | FLE_D.
+Scheme Equality for f_bin_op_D.
+Instance Decidable_eq_f_bin_op_D :
+forall (x y : f_bin_op_D), Decidable (x = y) :=
+Decidable_eq_from_dec f_bin_op_D_eq_dec.
+
+Inductive f_bin_op_S := FSGNJ_S | FSGNJN_S | FSGNJX_S | FMIN_S | FMAX_S | FEQ_S | FLT_S | FLE_S.
+Scheme Equality for f_bin_op_S.
+Instance Decidable_eq_f_bin_op_S :
+forall (x y : f_bin_op_S), Decidable (x = y) :=
+Decidable_eq_from_dec f_bin_op_S_eq_dec.
+
+Inductive f_bin_rm_op_D := FADD_D | FSUB_D | FMUL_D | FDIV_D.
+Scheme Equality for f_bin_rm_op_D.
+Instance Decidable_eq_f_bin_rm_op_D :
+forall (x y : f_bin_rm_op_D), Decidable (x = y) :=
+Decidable_eq_from_dec f_bin_rm_op_D_eq_dec.
+
+Inductive f_bin_rm_op_S := FADD_S | FSUB_S | FMUL_S | FDIV_S.
+Scheme Equality for f_bin_rm_op_S.
+Instance Decidable_eq_f_bin_rm_op_S :
+forall (x y : f_bin_rm_op_S), Decidable (x = y) :=
+Decidable_eq_from_dec f_bin_rm_op_S_eq_dec.
+
+Inductive f_madd_op_D := FMADD_D | FMSUB_D | FNMSUB_D | FNMADD_D.
+Scheme Equality for f_madd_op_D.
+Instance Decidable_eq_f_madd_op_D :
+forall (x y : f_madd_op_D), Decidable (x = y) :=
+Decidable_eq_from_dec f_madd_op_D_eq_dec.
+
+Inductive f_madd_op_S := FMADD_S | FMSUB_S | FNMSUB_S | FNMADD_S.
+Scheme Equality for f_madd_op_S.
+Instance Decidable_eq_f_madd_op_S :
+forall (x y : f_madd_op_S), Decidable (x = y) :=
+Decidable_eq_from_dec f_madd_op_S_eq_dec.
+
+Inductive f_un_op_D := FCLASS_D | FMV_X_D | FMV_D_X.
+Scheme Equality for f_un_op_D.
+Instance Decidable_eq_f_un_op_D :
+forall (x y : f_un_op_D), Decidable (x = y) :=
+Decidable_eq_from_dec f_un_op_D_eq_dec.
+
+Inductive f_un_op_S := FCLASS_S | FMV_X_W | FMV_W_X.
+Scheme Equality for f_un_op_S.
+Instance Decidable_eq_f_un_op_S :
+forall (x y : f_un_op_S), Decidable (x = y) :=
+Decidable_eq_from_dec f_un_op_S_eq_dec.
+
+Inductive f_un_rm_op_D :=
+ FSQRT_D
+ | FCVT_W_D
+ | FCVT_WU_D
+ | FCVT_D_W
+ | FCVT_D_WU
+ | FCVT_S_D
+ | FCVT_D_S
+ | FCVT_L_D
+ | FCVT_LU_D
+ | FCVT_D_L
+ | FCVT_D_LU.
+Scheme Equality for f_un_rm_op_D.
+Instance Decidable_eq_f_un_rm_op_D :
+forall (x y : f_un_rm_op_D), Decidable (x = y) :=
+Decidable_eq_from_dec f_un_rm_op_D_eq_dec.
+
+Inductive f_un_rm_op_S :=
+ FSQRT_S
+ | FCVT_W_S
+ | FCVT_WU_S
+ | FCVT_S_W
+ | FCVT_S_WU
+ | FCVT_L_S
+ | FCVT_LU_S
+ | FCVT_S_L
+ | FCVT_S_LU.
+Scheme Equality for f_un_rm_op_S.
+Instance Decidable_eq_f_un_rm_op_S :
+forall (x y : f_un_rm_op_S), Decidable (x = y) :=
+Decidable_eq_from_dec f_un_rm_op_S_eq_dec.
+
Inductive iop := RISCV_ADDI | RISCV_SLTI | RISCV_SLTIU | RISCV_XORI | RISCV_ORI | RISCV_ANDI.
Scheme Equality for iop.
Instance Decidable_eq_iop :
@@ -155,6 +281,12 @@ Instance Decidable_eq_ropw :
forall (x y : ropw), Decidable (x = y) :=
Decidable_eq_from_dec ropw_eq_dec.
+Inductive rounding_mode := RM_RNE | RM_RTZ | RM_RDN | RM_RUP | RM_RMM | RM_DYN.
+Scheme Equality for rounding_mode.
+Instance Decidable_eq_rounding_mode :
+forall (x y : rounding_mode), Decidable (x = y) :=
+Decidable_eq_from_dec rounding_mode_eq_dec.
+
Inductive sop := RISCV_SLLI | RISCV_SRLI | RISCV_SRAI.
Scheme Equality for sop.
Instance Decidable_eq_sop :
@@ -247,24 +379,46 @@ Inductive ast :=
| REMW : (regidx * regidx * regidx * bool) -> ast
| CSR : (bits 12 * regidx * regidx * bool * csrop) -> ast
| URET : unit -> ast
+ | LOAD_FP : (bits 12 * regidx * regidx * word_width) -> ast
+ | STORE_FP : (bits 12 * regidx * regidx * word_width) -> ast
+ | F_MADD_TYPE_S : (regidx * regidx * regidx * rounding_mode * regidx * f_madd_op_S) -> ast
+ | F_BIN_RM_TYPE_S : (regidx * regidx * rounding_mode * regidx * f_bin_rm_op_S) -> ast
+ | F_UN_RM_TYPE_S : (regidx * rounding_mode * regidx * f_un_rm_op_S) -> ast
+ | F_BIN_TYPE_S : (regidx * regidx * regidx * f_bin_op_S) -> ast
+ | F_UN_TYPE_S : (regidx * regidx * f_un_op_S) -> ast
+ | C_FLWSP : (bits 6 * regidx) -> ast
+ | C_FSWSP : (bits 6 * regidx) -> ast
+ | C_FLW : (bits 5 * cregidx * cregidx) -> ast
+ | C_FSW : (bits 5 * cregidx * cregidx) -> ast
+ | F_MADD_TYPE_D : (regidx * regidx * regidx * rounding_mode * regidx * f_madd_op_D) -> ast
+ | F_BIN_RM_TYPE_D : (regidx * regidx * rounding_mode * regidx * f_bin_rm_op_D) -> ast
+ | F_UN_RM_TYPE_D : (regidx * rounding_mode * regidx * f_un_rm_op_D) -> ast
+ | F_BIN_TYPE_D : (regidx * regidx * regidx * f_bin_op_D) -> ast
+ | F_UN_TYPE_D : (regidx * regidx * f_un_op_D) -> ast
+ | C_FLDSP : (bits 6 * regidx) -> ast
+ | C_FSDSP : (bits 6 * regidx) -> ast
+ | C_FLD : (bits 5 * cregidx * cregidx) -> ast
+ | C_FSD : (bits 5 * cregidx * cregidx) -> ast
| ILLEGAL : word -> ast
| C_ILLEGAL : half -> ast.
Arguments ast : clear implicits.
+Inductive PTW_Error :=
+ | PTW_Invalid_Addr : unit -> PTW_Error
+ | PTW_Access : unit -> PTW_Error
+ | PTW_Invalid_PTE : unit -> PTW_Error
+ | PTW_No_Permission : unit -> PTW_Error
+ | PTW_Misaligned : unit -> PTW_Error
+ | PTW_PTE_Update : unit -> PTW_Error
+ | PTW_Ext_Error : ext_ptw_error -> PTW_Error.
+Arguments PTW_Error : clear implicits.
+
Inductive Retired := RETIRE_SUCCESS | RETIRE_FAIL.
Scheme Equality for Retired.
Instance Decidable_eq_Retired :
forall (x y : Retired), Decidable (x = y) :=
Decidable_eq_from_dec Retired_eq_dec.
-Inductive AccessType := Read | Write | ReadWrite | Execute.
-Scheme Equality for AccessType.
-Instance Decidable_eq_AccessType :
-forall (x y : AccessType), Decidable (x = y) :=
-Decidable_eq_from_dec AccessType_eq_dec.
-
-Definition exc_code : Type := bits 8.
-
Inductive InterruptType :=
I_U_Software
| I_S_Software
@@ -280,29 +434,6 @@ Instance Decidable_eq_InterruptType :
forall (x y : InterruptType), Decidable (x = y) :=
Decidable_eq_from_dec InterruptType_eq_dec.
-Inductive ExceptionType :=
- E_Fetch_Addr_Align
- | E_Fetch_Access_Fault
- | E_Illegal_Instr
- | E_Breakpoint
- | E_Load_Addr_Align
- | E_Load_Access_Fault
- | E_SAMO_Addr_Align
- | E_SAMO_Access_Fault
- | E_U_EnvCall
- | E_S_EnvCall
- | E_Reserved_10
- | E_M_EnvCall
- | E_Fetch_Page_Fault
- | E_Load_Page_Fault
- | E_Reserved_14
- | E_SAMO_Page_Fault
- | E_CHERI.
-Scheme Equality for ExceptionType.
-Instance Decidable_eq_ExceptionType :
-forall (x y : ExceptionType), Decidable (x = y) :=
-Decidable_eq_from_dec ExceptionType_eq_dec.
-
Inductive exception :=
| Error_not_implemented : string -> exception | Error_internal_error : unit -> exception.
Arguments exception : clear implicits.
@@ -333,16 +464,54 @@ Decidable_eq_from_dec SATPMode_eq_dec.
Definition csrRW : Type := bits 2.
+Definition ext_access_type : Type := unit.
+
Definition regtype : Type := xlenbits.
+Definition fregtype : Type := flenbits.
+
Record Misa := { Misa_Misa_chunk_0 : mword 64; }.
Arguments Misa : clear implicits.
Notation "{[ r 'with' 'Misa_Misa_chunk_0' := e ]}" := {| Misa_Misa_chunk_0 := e |} (only parsing).
-Record SV48_PTE := { SV48_PTE_SV48_PTE_chunk_0 : mword 64; }.
-Arguments SV48_PTE : clear implicits.
-Notation "{[ r 'with' 'SV48_PTE_SV48_PTE_chunk_0' := e ]}" :=
- {| SV48_PTE_SV48_PTE_chunk_0 := e |} (only parsing).
+Record Counteren := { Counteren_Counteren_chunk_0 : mword 32; }.
+Arguments Counteren : clear implicits.
+Notation "{[ r 'with' 'Counteren_Counteren_chunk_0' := e ]}" :=
+ {| Counteren_Counteren_chunk_0 := e |} (only parsing).
+
+Record Counterin := { Counterin_Counterin_chunk_0 : mword 32; }.
+Arguments Counterin : clear implicits.
+Notation "{[ r 'with' 'Counterin_Counterin_chunk_0' := e ]}" :=
+ {| Counterin_Counterin_chunk_0 := e |} (only parsing).
+
+Record Fcsr := { Fcsr_Fcsr_chunk_0 : mword 32; }.
+Arguments Fcsr : clear implicits.
+Notation "{[ r 'with' 'Fcsr_Fcsr_chunk_0' := e ]}" := {| Fcsr_Fcsr_chunk_0 := e |} (only parsing).
+
+Record Mcause := { Mcause_Mcause_chunk_0 : mword 64; }.
+Arguments Mcause : clear implicits.
+Notation "{[ r 'with' 'Mcause_Mcause_chunk_0' := e ]}" :=
+ {| Mcause_Mcause_chunk_0 := e |} (only parsing).
+
+Record Medeleg := { Medeleg_Medeleg_chunk_0 : mword 64; }.
+Arguments Medeleg : clear implicits.
+Notation "{[ r 'with' 'Medeleg_Medeleg_chunk_0' := e ]}" :=
+ {| Medeleg_Medeleg_chunk_0 := e |} (only parsing).
+
+Record Minterrupts := { Minterrupts_Minterrupts_chunk_0 : mword 64; }.
+Arguments Minterrupts : clear implicits.
+Notation "{[ r 'with' 'Minterrupts_Minterrupts_chunk_0' := e ]}" :=
+ {| Minterrupts_Minterrupts_chunk_0 := e |} (only parsing).
+
+Record Mstatus := { Mstatus_Mstatus_chunk_0 : mword 64; }.
+Arguments Mstatus : clear implicits.
+Notation "{[ r 'with' 'Mstatus_Mstatus_chunk_0' := e ]}" :=
+ {| Mstatus_Mstatus_chunk_0 := e |} (only parsing).
+
+Record Mtvec := { Mtvec_Mtvec_chunk_0 : mword 64; }.
+Arguments Mtvec : clear implicits.
+Notation "{[ r 'with' 'Mtvec_Mtvec_chunk_0' := e ]}" :=
+ {| Mtvec_Mtvec_chunk_0 := e |} (only parsing).
Record PTE_Bits := { PTE_Bits_PTE_Bits_chunk_0 : mword 8; }.
Arguments PTE_Bits : clear implicits.
@@ -354,71 +523,91 @@ Arguments Pmpcfg_ent : clear implicits.
Notation "{[ r 'with' 'Pmpcfg_ent_Pmpcfg_ent_chunk_0' := e ]}" :=
{| Pmpcfg_ent_Pmpcfg_ent_chunk_0 := e |} (only parsing).
-Record Mstatus := { Mstatus_Mstatus_chunk_0 : mword 64; }.
-Arguments Mstatus : clear implicits.
-Notation "{[ r 'with' 'Mstatus_Mstatus_chunk_0' := e ]}" :=
- {| Mstatus_Mstatus_chunk_0 := e |} (only parsing).
+Record SV32_PTE := { SV32_PTE_SV32_PTE_chunk_0 : mword 32; }.
+Arguments SV32_PTE : clear implicits.
+Notation "{[ r 'with' 'SV32_PTE_SV32_PTE_chunk_0' := e ]}" :=
+ {| SV32_PTE_SV32_PTE_chunk_0 := e |} (only parsing).
-Record Sstatus := { Sstatus_Sstatus_chunk_0 : mword 64; }.
-Arguments Sstatus : clear implicits.
-Notation "{[ r 'with' 'Sstatus_Sstatus_chunk_0' := e ]}" :=
- {| Sstatus_Sstatus_chunk_0 := e |} (only parsing).
+Record SV32_Paddr := { SV32_Paddr_SV32_Paddr_chunk_0 : mword 34; }.
+Arguments SV32_Paddr : clear implicits.
+Notation "{[ r 'with' 'SV32_Paddr_SV32_Paddr_chunk_0' := e ]}" :=
+ {| SV32_Paddr_SV32_Paddr_chunk_0 := e |} (only parsing).
-Record Ustatus := { Ustatus_Ustatus_chunk_0 : mword 64; }.
-Arguments Ustatus : clear implicits.
-Notation "{[ r 'with' 'Ustatus_Ustatus_chunk_0' := e ]}" :=
- {| Ustatus_Ustatus_chunk_0 := e |} (only parsing).
+Record SV32_Vaddr := { SV32_Vaddr_SV32_Vaddr_chunk_0 : mword 32; }.
+Arguments SV32_Vaddr : clear implicits.
+Notation "{[ r 'with' 'SV32_Vaddr_SV32_Vaddr_chunk_0' := e ]}" :=
+ {| SV32_Vaddr_SV32_Vaddr_chunk_0 := e |} (only parsing).
-Record Minterrupts := { Minterrupts_Minterrupts_chunk_0 : mword 64; }.
-Arguments Minterrupts : clear implicits.
-Notation "{[ r 'with' 'Minterrupts_Minterrupts_chunk_0' := e ]}" :=
- {| Minterrupts_Minterrupts_chunk_0 := e |} (only parsing).
+Record SV39_PTE := { SV39_PTE_SV39_PTE_chunk_0 : mword 64; }.
+Arguments SV39_PTE : clear implicits.
+Notation "{[ r 'with' 'SV39_PTE_SV39_PTE_chunk_0' := e ]}" :=
+ {| SV39_PTE_SV39_PTE_chunk_0 := e |} (only parsing).
-Record Sinterrupts := { Sinterrupts_Sinterrupts_chunk_0 : mword 64; }.
-Arguments Sinterrupts : clear implicits.
-Notation "{[ r 'with' 'Sinterrupts_Sinterrupts_chunk_0' := e ]}" :=
- {| Sinterrupts_Sinterrupts_chunk_0 := e |} (only parsing).
+Record SV39_Paddr := { SV39_Paddr_SV39_Paddr_chunk_0 : mword 56; }.
+Arguments SV39_Paddr : clear implicits.
+Notation "{[ r 'with' 'SV39_Paddr_SV39_Paddr_chunk_0' := e ]}" :=
+ {| SV39_Paddr_SV39_Paddr_chunk_0 := e |} (only parsing).
-Record Uinterrupts := { Uinterrupts_Uinterrupts_chunk_0 : mword 64; }.
-Arguments Uinterrupts : clear implicits.
-Notation "{[ r 'with' 'Uinterrupts_Uinterrupts_chunk_0' := e ]}" :=
- {| Uinterrupts_Uinterrupts_chunk_0 := e |} (only parsing).
+Record SV39_Vaddr := { SV39_Vaddr_SV39_Vaddr_chunk_0 : mword 39; }.
+Arguments SV39_Vaddr : clear implicits.
+Notation "{[ r 'with' 'SV39_Vaddr_SV39_Vaddr_chunk_0' := e ]}" :=
+ {| SV39_Vaddr_SV39_Vaddr_chunk_0 := e |} (only parsing).
-Record Medeleg := { Medeleg_Medeleg_chunk_0 : mword 64; }.
-Arguments Medeleg : clear implicits.
-Notation "{[ r 'with' 'Medeleg_Medeleg_chunk_0' := e ]}" :=
- {| Medeleg_Medeleg_chunk_0 := e |} (only parsing).
+Record SV48_PTE := { SV48_PTE_SV48_PTE_chunk_0 : mword 64; }.
+Arguments SV48_PTE : clear implicits.
+Notation "{[ r 'with' 'SV48_PTE_SV48_PTE_chunk_0' := e ]}" :=
+ {| SV48_PTE_SV48_PTE_chunk_0 := e |} (only parsing).
-Record Sedeleg := { Sedeleg_Sedeleg_chunk_0 : mword 64; }.
-Arguments Sedeleg : clear implicits.
-Notation "{[ r 'with' 'Sedeleg_Sedeleg_chunk_0' := e ]}" :=
- {| Sedeleg_Sedeleg_chunk_0 := e |} (only parsing).
+Record SV48_Paddr := { SV48_Paddr_SV48_Paddr_chunk_0 : mword 56; }.
+Arguments SV48_Paddr : clear implicits.
+Notation "{[ r 'with' 'SV48_Paddr_SV48_Paddr_chunk_0' := e ]}" :=
+ {| SV48_Paddr_SV48_Paddr_chunk_0 := e |} (only parsing).
-Record Mtvec := { Mtvec_Mtvec_chunk_0 : mword 64; }.
-Arguments Mtvec : clear implicits.
-Notation "{[ r 'with' 'Mtvec_Mtvec_chunk_0' := e ]}" :=
- {| Mtvec_Mtvec_chunk_0 := e |} (only parsing).
+Record SV48_Vaddr := { SV48_Vaddr_SV48_Vaddr_chunk_0 : mword 48; }.
+Arguments SV48_Vaddr : clear implicits.
+Notation "{[ r 'with' 'SV48_Vaddr_SV48_Vaddr_chunk_0' := e ]}" :=
+ {| SV48_Vaddr_SV48_Vaddr_chunk_0 := e |} (only parsing).
Record Satp32 := { Satp32_Satp32_chunk_0 : mword 32; }.
Arguments Satp32 : clear implicits.
Notation "{[ r 'with' 'Satp32_Satp32_chunk_0' := e ]}" :=
{| Satp32_Satp32_chunk_0 := e |} (only parsing).
-Record Mcause := { Mcause_Mcause_chunk_0 : mword 64; }.
-Arguments Mcause : clear implicits.
-Notation "{[ r 'with' 'Mcause_Mcause_chunk_0' := e ]}" :=
- {| Mcause_Mcause_chunk_0 := e |} (only parsing).
-
-Record Counteren := { Counteren_Counteren_chunk_0 : mword 32; }.
-Arguments Counteren : clear implicits.
-Notation "{[ r 'with' 'Counteren_Counteren_chunk_0' := e ]}" :=
- {| Counteren_Counteren_chunk_0 := e |} (only parsing).
-
Record Satp64 := { Satp64_Satp64_chunk_0 : mword 64; }.
Arguments Satp64 : clear implicits.
Notation "{[ r 'with' 'Satp64_Satp64_chunk_0' := e ]}" :=
{| Satp64_Satp64_chunk_0 := e |} (only parsing).
+Record Sedeleg := { Sedeleg_Sedeleg_chunk_0 : mword 64; }.
+Arguments Sedeleg : clear implicits.
+Notation "{[ r 'with' 'Sedeleg_Sedeleg_chunk_0' := e ]}" :=
+ {| Sedeleg_Sedeleg_chunk_0 := e |} (only parsing).
+
+Record Sinterrupts := { Sinterrupts_Sinterrupts_chunk_0 : mword 64; }.
+Arguments Sinterrupts : clear implicits.
+Notation "{[ r 'with' 'Sinterrupts_Sinterrupts_chunk_0' := e ]}" :=
+ {| Sinterrupts_Sinterrupts_chunk_0 := e |} (only parsing).
+
+Record Sstatus := { Sstatus_Sstatus_chunk_0 : mword 64; }.
+Arguments Sstatus : clear implicits.
+Notation "{[ r 'with' 'Sstatus_Sstatus_chunk_0' := e ]}" :=
+ {| Sstatus_Sstatus_chunk_0 := e |} (only parsing).
+
+Record Uinterrupts := { Uinterrupts_Uinterrupts_chunk_0 : mword 64; }.
+Arguments Uinterrupts : clear implicits.
+Notation "{[ r 'with' 'Uinterrupts_Uinterrupts_chunk_0' := e ]}" :=
+ {| Uinterrupts_Uinterrupts_chunk_0 := e |} (only parsing).
+
+Record Ustatus := { Ustatus_Ustatus_chunk_0 : mword 64; }.
+Arguments Ustatus : clear implicits.
+Notation "{[ r 'with' 'Ustatus_Ustatus_chunk_0' := e ]}" :=
+ {| Ustatus_Ustatus_chunk_0 := e |} (only parsing).
+
+Record htif_cmd := { htif_cmd_htif_cmd_chunk_0 : mword 64; }.
+Arguments htif_cmd : clear implicits.
+Notation "{[ r 'with' 'htif_cmd_htif_cmd_chunk_0' := e ]}" :=
+ {| htif_cmd_htif_cmd_chunk_0 := e |} (only parsing).
+
Inductive PmpAddrMatchType := OFF | TOR | NA4 | NAPOT.
Scheme Equality for PmpAddrMatchType.
Instance Decidable_eq_PmpAddrMatchType :
@@ -466,21 +655,28 @@ Record sync_exception :=
sync_exception_excinfo : option xlenbits;
sync_exception_ext : option ext_exception; }.
Arguments sync_exception : clear implicits.
-Notation "{[ r 'with' 'sync_exception_trap' := e ]}" := {|
- sync_exception_trap := e;
- sync_exception_excinfo := sync_exception_excinfo r;
- sync_exception_ext := sync_exception_ext r
- |}.
-Notation "{[ r 'with' 'sync_exception_excinfo' := e ]}" := {|
- sync_exception_excinfo := e;
- sync_exception_trap := sync_exception_trap r;
- sync_exception_ext := sync_exception_ext r
- |}.
-Notation "{[ r 'with' 'sync_exception_ext' := e ]}" := {|
- sync_exception_ext := e;
- sync_exception_trap := sync_exception_trap r;
- sync_exception_excinfo := sync_exception_excinfo r
- |}.
+Notation "{[ r 'with' 'sync_exception_trap' := e ]}" :=
+ match r with Build_sync_exception _ f1 f2 => Build_sync_exception e f1 f2 end.
+Notation "{[ r 'with' 'sync_exception_excinfo' := e ]}" :=
+ match r with Build_sync_exception f0 _ f2 => Build_sync_exception f0 e f2 end.
+Notation "{[ r 'with' 'sync_exception_ext' := e ]}" :=
+ match r with Build_sync_exception f0 f1 _ => Build_sync_exception f0 f1 e end.
+
+Definition bits_rm : Type := bits 3.
+
+Definition bits_fflags : Type := bits 5.
+
+Definition bits_S : Type := bits 32.
+
+Definition bits_D : Type := bits 64.
+
+Definition bits_W : Type := bits 32.
+
+Definition bits_WU : Type := bits 32.
+
+Definition bits_L : Type := bits 64.
+
+Definition bits_LU : Type := bits 64.
Inductive interrupt_set :=
| Ints_Pending : xlenbits -> interrupt_set
@@ -499,19 +695,13 @@ Inductive MemoryOpResult {a : Type} :=
| MemValue : a -> MemoryOpResult | MemException : ExceptionType -> MemoryOpResult.
Arguments MemoryOpResult : clear implicits.
-Record htif_cmd := { htif_cmd_htif_cmd_chunk_0 : mword 64; }.
-Arguments htif_cmd : clear implicits.
-Notation "{[ r 'with' 'htif_cmd_htif_cmd_chunk_0' := e ]}" :=
- {| htif_cmd_htif_cmd_chunk_0 := e |} (only parsing).
-
Definition pteAttribs : Type := bits 8.
-Inductive PTW_Error :=
- PTW_Access | PTW_Invalid_PTE | PTW_No_Permission | PTW_Misaligned | PTW_PTE_Update.
-Scheme Equality for PTW_Error.
-Instance Decidable_eq_PTW_Error :
-forall (x y : PTW_Error), Decidable (x = y) :=
-Decidable_eq_from_dec PTW_Error_eq_dec.
+Definition extPte : Type := bits 10.
+
+Inductive PTE_Check :=
+ | PTE_Check_Success : ext_ptw -> PTE_Check | PTE_Check_Failure : ext_ptw -> PTE_Check.
+Arguments PTE_Check : clear implicits.
Definition vaddr32 : Type := bits 32.
@@ -521,31 +711,6 @@ Definition pte32 : Type := bits 32.
Definition asid32 : Type := bits 9.
-Record SV32_Vaddr := { SV32_Vaddr_SV32_Vaddr_chunk_0 : mword 32; }.
-Arguments SV32_Vaddr : clear implicits.
-Notation "{[ r 'with' 'SV32_Vaddr_SV32_Vaddr_chunk_0' := e ]}" :=
- {| SV32_Vaddr_SV32_Vaddr_chunk_0 := e |} (only parsing).
-
-Record SV48_Vaddr := { SV48_Vaddr_SV48_Vaddr_chunk_0 : mword 48; }.
-Arguments SV48_Vaddr : clear implicits.
-Notation "{[ r 'with' 'SV48_Vaddr_SV48_Vaddr_chunk_0' := e ]}" :=
- {| SV48_Vaddr_SV48_Vaddr_chunk_0 := e |} (only parsing).
-
-Record SV48_Paddr := { SV48_Paddr_SV48_Paddr_chunk_0 : mword 56; }.
-Arguments SV48_Paddr : clear implicits.
-Notation "{[ r 'with' 'SV48_Paddr_SV48_Paddr_chunk_0' := e ]}" :=
- {| SV48_Paddr_SV48_Paddr_chunk_0 := e |} (only parsing).
-
-Record SV32_Paddr := { SV32_Paddr_SV32_Paddr_chunk_0 : mword 34; }.
-Arguments SV32_Paddr : clear implicits.
-Notation "{[ r 'with' 'SV32_Paddr_SV32_Paddr_chunk_0' := e ]}" :=
- {| SV32_Paddr_SV32_Paddr_chunk_0 := e |} (only parsing).
-
-Record SV32_PTE := { SV32_PTE_SV32_PTE_chunk_0 : mword 32; }.
-Arguments SV32_PTE : clear implicits.
-Notation "{[ r 'with' 'SV32_PTE_SV32_PTE_chunk_0' := e ]}" :=
- {| SV32_PTE_SV32_PTE_chunk_0 := e |} (only parsing).
-
Definition paddr64 : Type := bits 56.
Definition pte64 : Type := bits 64.
@@ -554,32 +719,17 @@ Definition asid64 : Type := bits 16.
Definition vaddr39 : Type := bits 39.
-Record SV39_Vaddr := { SV39_Vaddr_SV39_Vaddr_chunk_0 : mword 39; }.
-Arguments SV39_Vaddr : clear implicits.
-Notation "{[ r 'with' 'SV39_Vaddr_SV39_Vaddr_chunk_0' := e ]}" :=
- {| SV39_Vaddr_SV39_Vaddr_chunk_0 := e |} (only parsing).
-
-Record SV39_Paddr := { SV39_Paddr_SV39_Paddr_chunk_0 : mword 56; }.
-Arguments SV39_Paddr : clear implicits.
-Notation "{[ r 'with' 'SV39_Paddr_SV39_Paddr_chunk_0' := e ]}" :=
- {| SV39_Paddr_SV39_Paddr_chunk_0 := e |} (only parsing).
-
-Record SV39_PTE := { SV39_PTE_SV39_PTE_chunk_0 : mword 64; }.
-Arguments SV39_PTE : clear implicits.
-Notation "{[ r 'with' 'SV39_PTE_SV39_PTE_chunk_0' := e ]}" :=
- {| SV39_PTE_SV39_PTE_chunk_0 := e |} (only parsing).
-
Definition vaddr48 : Type := bits 48.
Definition pte48 : Type := bits 64.
Inductive PTW_Result {paddr : Type} {pte : Type} :=
- | PTW_Success : (paddr * pte * paddr * {n : Z & ArithFact (n >= 0)} * bool) -> PTW_Result
- | PTW_Failure : PTW_Error -> PTW_Result.
+ | PTW_Success : (paddr * pte * paddr * {n : Z & ArithFact (n >=? 0)} * bool * ext_ptw) -> PTW_Result
+ | PTW_Failure : (PTW_Error * ext_ptw) -> PTW_Result.
Arguments PTW_Result : clear implicits.
Inductive TR_Result {paddr : Type} {failure : Type} :=
- | TR_Address : paddr -> TR_Result | TR_Failure : failure -> TR_Result.
+ | TR_Address : (paddr * ext_ptw) -> TR_Result | TR_Failure : (failure * ext_ptw) -> TR_Result.
Arguments TR_Result : clear implicits.
Record TLB_Entry {asidlen : Z} {valen : Z} {palen : Z} {ptelen : Z} :=
@@ -593,115 +743,45 @@ Record TLB_Entry {asidlen : Z} {valen : Z} {palen : Z} {ptelen : Z} :=
TLB_Entry_pteAddr : bits palen;
TLB_Entry_age : bits 64; }.
Arguments TLB_Entry : clear implicits.
-Notation "{[ r 'with' 'TLB_Entry_asid' := e ]}" := {|
- TLB_Entry_asid := e;
- TLB_Entry_global := TLB_Entry_global r;
- TLB_Entry_vAddr := TLB_Entry_vAddr r;
- TLB_Entry_pAddr := TLB_Entry_pAddr r;
- TLB_Entry_vMatchMask := TLB_Entry_vMatchMask r;
- TLB_Entry_vAddrMask := TLB_Entry_vAddrMask r;
- TLB_Entry_pte := TLB_Entry_pte r;
- TLB_Entry_pteAddr := TLB_Entry_pteAddr r;
- TLB_Entry_age := TLB_Entry_age r
- |}.
-Notation "{[ r 'with' 'TLB_Entry_global' := e ]}" := {|
- TLB_Entry_global := e;
- TLB_Entry_asid := TLB_Entry_asid r;
- TLB_Entry_vAddr := TLB_Entry_vAddr r;
- TLB_Entry_pAddr := TLB_Entry_pAddr r;
- TLB_Entry_vMatchMask := TLB_Entry_vMatchMask r;
- TLB_Entry_vAddrMask := TLB_Entry_vAddrMask r;
- TLB_Entry_pte := TLB_Entry_pte r;
- TLB_Entry_pteAddr := TLB_Entry_pteAddr r;
- TLB_Entry_age := TLB_Entry_age r
- |}.
-Notation "{[ r 'with' 'TLB_Entry_vAddr' := e ]}" := {|
- TLB_Entry_vAddr := e;
- TLB_Entry_asid := TLB_Entry_asid r;
- TLB_Entry_global := TLB_Entry_global r;
- TLB_Entry_pAddr := TLB_Entry_pAddr r;
- TLB_Entry_vMatchMask := TLB_Entry_vMatchMask r;
- TLB_Entry_vAddrMask := TLB_Entry_vAddrMask r;
- TLB_Entry_pte := TLB_Entry_pte r;
- TLB_Entry_pteAddr := TLB_Entry_pteAddr r;
- TLB_Entry_age := TLB_Entry_age r
- |}.
-Notation "{[ r 'with' 'TLB_Entry_pAddr' := e ]}" := {|
- TLB_Entry_pAddr := e;
- TLB_Entry_asid := TLB_Entry_asid r;
- TLB_Entry_global := TLB_Entry_global r;
- TLB_Entry_vAddr := TLB_Entry_vAddr r;
- TLB_Entry_vMatchMask := TLB_Entry_vMatchMask r;
- TLB_Entry_vAddrMask := TLB_Entry_vAddrMask r;
- TLB_Entry_pte := TLB_Entry_pte r;
- TLB_Entry_pteAddr := TLB_Entry_pteAddr r;
- TLB_Entry_age := TLB_Entry_age r
- |}.
-Notation "{[ r 'with' 'TLB_Entry_vMatchMask' := e ]}" := {|
- TLB_Entry_vMatchMask := e;
- TLB_Entry_asid := TLB_Entry_asid r;
- TLB_Entry_global := TLB_Entry_global r;
- TLB_Entry_vAddr := TLB_Entry_vAddr r;
- TLB_Entry_pAddr := TLB_Entry_pAddr r;
- TLB_Entry_vAddrMask := TLB_Entry_vAddrMask r;
- TLB_Entry_pte := TLB_Entry_pte r;
- TLB_Entry_pteAddr := TLB_Entry_pteAddr r;
- TLB_Entry_age := TLB_Entry_age r
- |}.
-Notation "{[ r 'with' 'TLB_Entry_vAddrMask' := e ]}" := {|
- TLB_Entry_vAddrMask := e;
- TLB_Entry_asid := TLB_Entry_asid r;
- TLB_Entry_global := TLB_Entry_global r;
- TLB_Entry_vAddr := TLB_Entry_vAddr r;
- TLB_Entry_pAddr := TLB_Entry_pAddr r;
- TLB_Entry_vMatchMask := TLB_Entry_vMatchMask r;
- TLB_Entry_pte := TLB_Entry_pte r;
- TLB_Entry_pteAddr := TLB_Entry_pteAddr r;
- TLB_Entry_age := TLB_Entry_age r
- |}.
-Notation "{[ r 'with' 'TLB_Entry_pte' := e ]}" := {|
- TLB_Entry_pte := e;
- TLB_Entry_asid := TLB_Entry_asid r;
- TLB_Entry_global := TLB_Entry_global r;
- TLB_Entry_vAddr := TLB_Entry_vAddr r;
- TLB_Entry_pAddr := TLB_Entry_pAddr r;
- TLB_Entry_vMatchMask := TLB_Entry_vMatchMask r;
- TLB_Entry_vAddrMask := TLB_Entry_vAddrMask r;
- TLB_Entry_pteAddr := TLB_Entry_pteAddr r;
- TLB_Entry_age := TLB_Entry_age r
- |}.
-Notation "{[ r 'with' 'TLB_Entry_pteAddr' := e ]}" := {|
- TLB_Entry_pteAddr := e;
- TLB_Entry_asid := TLB_Entry_asid r;
- TLB_Entry_global := TLB_Entry_global r;
- TLB_Entry_vAddr := TLB_Entry_vAddr r;
- TLB_Entry_pAddr := TLB_Entry_pAddr r;
- TLB_Entry_vMatchMask := TLB_Entry_vMatchMask r;
- TLB_Entry_vAddrMask := TLB_Entry_vAddrMask r;
- TLB_Entry_pte := TLB_Entry_pte r;
- TLB_Entry_age := TLB_Entry_age r
- |}.
-Notation "{[ r 'with' 'TLB_Entry_age' := e ]}" := {|
- TLB_Entry_age := e;
- TLB_Entry_asid := TLB_Entry_asid r;
- TLB_Entry_global := TLB_Entry_global r;
- TLB_Entry_vAddr := TLB_Entry_vAddr r;
- TLB_Entry_pAddr := TLB_Entry_pAddr r;
- TLB_Entry_vMatchMask := TLB_Entry_vMatchMask r;
- TLB_Entry_vAddrMask := TLB_Entry_vAddrMask r;
- TLB_Entry_pte := TLB_Entry_pte r;
- TLB_Entry_pteAddr := TLB_Entry_pteAddr r
- |}.
+Notation "{[ r 'with' 'TLB_Entry_asid' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ _ f1 f2 f3 f4 f5 f6 f7 f8 =>
+ Build_TLB_Entry _ _ _ _ e f1 f2 f3 f4 f5 f6 f7 f8 end.
+Notation "{[ r 'with' 'TLB_Entry_global' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ f0 _ f2 f3 f4 f5 f6 f7 f8 =>
+ Build_TLB_Entry _ _ _ _ f0 e f2 f3 f4 f5 f6 f7 f8 end.
+Notation "{[ r 'with' 'TLB_Entry_vAddr' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ f0 f1 _ f3 f4 f5 f6 f7 f8 =>
+ Build_TLB_Entry _ _ _ _ f0 f1 e f3 f4 f5 f6 f7 f8 end.
+Notation "{[ r 'with' 'TLB_Entry_pAddr' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ f0 f1 f2 _ f4 f5 f6 f7 f8 =>
+ Build_TLB_Entry _ _ _ _ f0 f1 f2 e f4 f5 f6 f7 f8 end.
+Notation "{[ r 'with' 'TLB_Entry_vMatchMask' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 _ f5 f6 f7 f8 =>
+ Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 e f5 f6 f7 f8 end.
+Notation "{[ r 'with' 'TLB_Entry_vAddrMask' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 f4 _ f6 f7 f8 =>
+ Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 f4 e f6 f7 f8 end.
+Notation "{[ r 'with' 'TLB_Entry_pte' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 f4 f5 _ f7 f8 =>
+ Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 f4 f5 e f7 f8 end.
+Notation "{[ r 'with' 'TLB_Entry_pteAddr' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 f4 f5 f6 _ f8 =>
+ Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 f4 f5 f6 e f8 end.
+Notation "{[ r 'with' 'TLB_Entry_age' := e ]}" :=
+ match r with Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 f4 f5 f6 f7 _ =>
+ Build_TLB_Entry _ _ _ _ f0 f1 f2 f3 f4 f5 f6 f7 e end.
Definition TLB39_Entry : Type := TLB_Entry 16 39 56 64.
Definition TLB48_Entry : Type := TLB_Entry 16 48 56 64.
Inductive register_value :=
- | Regval_vector : (Z * bool * list register_value) -> register_value
+ | Regval_vector : list register_value -> register_value
| Regval_list : list register_value -> register_value
| Regval_option : option register_value -> register_value
| Regval_Counteren : Counteren -> register_value
+ | Regval_Counterin : Counterin -> register_value
+ | Regval_Fcsr : Fcsr -> register_value
| Regval_Mcause : Mcause -> register_value
| Regval_Medeleg : Medeleg -> register_value
| Regval_Minterrupts : Minterrupts -> register_value
@@ -728,6 +808,42 @@ Record regstate :=
htif_done : bool;
htif_tohost : mword 64;
mtimecmp : mword 64;
+ fcsr : Fcsr;
+ f31 : mword 64;
+ f30 : mword 64;
+ f29 : mword 64;
+ f28 : mword 64;
+ f27 : mword 64;
+ f26 : mword 64;
+ f25 : mword 64;
+ f24 : mword 64;
+ f23 : mword 64;
+ f22 : mword 64;
+ f21 : mword 64;
+ f20 : mword 64;
+ f19 : mword 64;
+ f18 : mword 64;
+ f17 : mword 64;
+ f16 : mword 64;
+ f15 : mword 64;
+ f14 : mword 64;
+ f13 : mword 64;
+ f12 : mword 64;
+ f11 : mword 64;
+ f10 : mword 64;
+ f9 : mword 64;
+ f8 : mword 64;
+ f7 : mword 64;
+ f6 : mword 64;
+ f5 : mword 64;
+ f4 : mword 64;
+ f3 : mword 64;
+ f2 : mword 64;
+ f1 : mword 64;
+ f0 : mword 64;
+ fs : vec (mword 64) 32;
+ float_fflags : mword 64;
+ float_result : mword 64;
utval : mword 64;
ucause : Mcause;
uepc : mword 64;
@@ -781,6 +897,7 @@ Record regstate :=
minstret : mword 64;
mtime : mword 64;
mcycle : mword 64;
+ mcountinhibit : Counterin;
scounteren : Counteren;
mcounteren : Counteren;
mscratch : mword 64;
@@ -832,12472 +949,704 @@ Record regstate :=
nextPC : mword 64;
PC : mword 64; }.
Arguments regstate : clear implicits.
-Notation "{[ r 'with' 'satp' := e ]}" := {|
- satp := e;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'tlb48' := e ]}" := {|
- tlb48 := e;
- satp := satp r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'tlb39' := e ]}" := {|
- tlb39 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'htif_exit_code' := e ]}" := {|
- htif_exit_code := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'htif_done' := e ]}" := {|
- htif_done := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'htif_tohost' := e ]}" := {|
- htif_tohost := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mtimecmp' := e ]}" := {|
- mtimecmp := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'utval' := e ]}" := {|
- utval := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'ucause' := e ]}" := {|
- ucause := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'uepc' := e ]}" := {|
- uepc := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'uscratch' := e ]}" := {|
- uscratch := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'utvec' := e ]}" := {|
- utvec := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr15' := e ]}" := {|
- pmpaddr15 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr14' := e ]}" := {|
- pmpaddr14 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr13' := e ]}" := {|
- pmpaddr13 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr12' := e ]}" := {|
- pmpaddr12 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr11' := e ]}" := {|
- pmpaddr11 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr10' := e ]}" := {|
- pmpaddr10 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr9' := e ]}" := {|
- pmpaddr9 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr8' := e ]}" := {|
- pmpaddr8 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr7' := e ]}" := {|
- pmpaddr7 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr6' := e ]}" := {|
- pmpaddr6 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr5' := e ]}" := {|
- pmpaddr5 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr4' := e ]}" := {|
- pmpaddr4 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr3' := e ]}" := {|
- pmpaddr3 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr2' := e ]}" := {|
- pmpaddr2 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr1' := e ]}" := {|
- pmpaddr1 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmpaddr0' := e ]}" := {|
- pmpaddr0 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp15cfg' := e ]}" := {|
- pmp15cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp14cfg' := e ]}" := {|
- pmp14cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp13cfg' := e ]}" := {|
- pmp13cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp12cfg' := e ]}" := {|
- pmp12cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp11cfg' := e ]}" := {|
- pmp11cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp10cfg' := e ]}" := {|
- pmp10cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp9cfg' := e ]}" := {|
- pmp9cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp8cfg' := e ]}" := {|
- pmp8cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp7cfg' := e ]}" := {|
- pmp7cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp6cfg' := e ]}" := {|
- pmp6cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp5cfg' := e ]}" := {|
- pmp5cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp4cfg' := e ]}" := {|
- pmp4cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp3cfg' := e ]}" := {|
- pmp3cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp2cfg' := e ]}" := {|
- pmp2cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp1cfg' := e ]}" := {|
- pmp1cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'pmp0cfg' := e ]}" := {|
- pmp0cfg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'tselect' := e ]}" := {|
- tselect := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'stval' := e ]}" := {|
- stval := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'scause' := e ]}" := {|
- scause := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'sepc' := e ]}" := {|
- sepc := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'sscratch' := e ]}" := {|
- sscratch := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'stvec' := e ]}" := {|
- stvec := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'sideleg' := e ]}" := {|
- sideleg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'sedeleg' := e ]}" := {|
- sedeleg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mhartid' := e ]}" := {|
- mhartid := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'marchid' := e ]}" := {|
- marchid := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mimpid' := e ]}" := {|
- mimpid := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mvendorid' := e ]}" := {|
- mvendorid := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'minstret_written' := e ]}" := {|
- minstret_written := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'minstret' := e ]}" := {|
- minstret := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mtime' := e ]}" := {|
- mtime := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mcycle' := e ]}" := {|
- mcycle := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'scounteren' := e ]}" := {|
- scounteren := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mcounteren' := e ]}" := {|
- mcounteren := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mscratch' := e ]}" := {|
- mscratch := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mtval' := e ]}" := {|
- mtval := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mepc' := e ]}" := {|
- mepc := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mcause' := e ]}" := {|
- mcause := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mtvec' := e ]}" := {|
- mtvec := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'medeleg' := e ]}" := {|
- medeleg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mideleg' := e ]}" := {|
- mideleg := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mie' := e ]}" := {|
- mie := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mip' := e ]}" := {|
- mip := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'mstatus' := e ]}" := {|
- mstatus := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'misa' := e ]}" := {|
- misa := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'cur_inst' := e ]}" := {|
- cur_inst := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'cur_privilege' := e ]}" := {|
- cur_privilege := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x31' := e ]}" := {|
- x31 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x30' := e ]}" := {|
- x30 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x29' := e ]}" := {|
- x29 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x28' := e ]}" := {|
- x28 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x27' := e ]}" := {|
- x27 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x26' := e ]}" := {|
- x26 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x25' := e ]}" := {|
- x25 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x24' := e ]}" := {|
- x24 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x23' := e ]}" := {|
- x23 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x22' := e ]}" := {|
- x22 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x21' := e ]}" := {|
- x21 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x20' := e ]}" := {|
- x20 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x19' := e ]}" := {|
- x19 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x18' := e ]}" := {|
- x18 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x17' := e ]}" := {|
- x17 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x16' := e ]}" := {|
- x16 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x15' := e ]}" := {|
- x15 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x14' := e ]}" := {|
- x14 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x13' := e ]}" := {|
- x13 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x12' := e ]}" := {|
- x12 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x11' := e ]}" := {|
- x11 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x10' := e ]}" := {|
- x10 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x9' := e ]}" := {|
- x9 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x8' := e ]}" := {|
- x8 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x7' := e ]}" := {|
- x7 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x6' := e ]}" := {|
- x6 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x5' := e ]}" := {|
- x5 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x4' := e ]}" := {|
- x4 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x3' := e ]}" := {|
- x3 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x2' := e ]}" := {|
- x2 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'x1' := e ]}" := {|
- x1 := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'Xs' := e ]}" := {|
- Xs := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- instbits := instbits r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'instbits' := e ]}" := {|
- instbits := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- nextPC := nextPC r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'nextPC' := e ]}" := {|
- nextPC := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- PC := PC r
- |}.
-Notation "{[ r 'with' 'PC' := e ]}" := {|
- PC := e;
- satp := satp r;
- tlb48 := tlb48 r;
- tlb39 := tlb39 r;
- htif_exit_code := htif_exit_code r;
- htif_done := htif_done r;
- htif_tohost := htif_tohost r;
- mtimecmp := mtimecmp r;
- utval := utval r;
- ucause := ucause r;
- uepc := uepc r;
- uscratch := uscratch r;
- utvec := utvec r;
- pmpaddr15 := pmpaddr15 r;
- pmpaddr14 := pmpaddr14 r;
- pmpaddr13 := pmpaddr13 r;
- pmpaddr12 := pmpaddr12 r;
- pmpaddr11 := pmpaddr11 r;
- pmpaddr10 := pmpaddr10 r;
- pmpaddr9 := pmpaddr9 r;
- pmpaddr8 := pmpaddr8 r;
- pmpaddr7 := pmpaddr7 r;
- pmpaddr6 := pmpaddr6 r;
- pmpaddr5 := pmpaddr5 r;
- pmpaddr4 := pmpaddr4 r;
- pmpaddr3 := pmpaddr3 r;
- pmpaddr2 := pmpaddr2 r;
- pmpaddr1 := pmpaddr1 r;
- pmpaddr0 := pmpaddr0 r;
- pmp15cfg := pmp15cfg r;
- pmp14cfg := pmp14cfg r;
- pmp13cfg := pmp13cfg r;
- pmp12cfg := pmp12cfg r;
- pmp11cfg := pmp11cfg r;
- pmp10cfg := pmp10cfg r;
- pmp9cfg := pmp9cfg r;
- pmp8cfg := pmp8cfg r;
- pmp7cfg := pmp7cfg r;
- pmp6cfg := pmp6cfg r;
- pmp5cfg := pmp5cfg r;
- pmp4cfg := pmp4cfg r;
- pmp3cfg := pmp3cfg r;
- pmp2cfg := pmp2cfg r;
- pmp1cfg := pmp1cfg r;
- pmp0cfg := pmp0cfg r;
- tselect := tselect r;
- stval := stval r;
- scause := scause r;
- sepc := sepc r;
- sscratch := sscratch r;
- stvec := stvec r;
- sideleg := sideleg r;
- sedeleg := sedeleg r;
- mhartid := mhartid r;
- marchid := marchid r;
- mimpid := mimpid r;
- mvendorid := mvendorid r;
- minstret_written := minstret_written r;
- minstret := minstret r;
- mtime := mtime r;
- mcycle := mcycle r;
- scounteren := scounteren r;
- mcounteren := mcounteren r;
- mscratch := mscratch r;
- mtval := mtval r;
- mepc := mepc r;
- mcause := mcause r;
- mtvec := mtvec r;
- medeleg := medeleg r;
- mideleg := mideleg r;
- mie := mie r;
- mip := mip r;
- mstatus := mstatus r;
- misa := misa r;
- cur_inst := cur_inst r;
- cur_privilege := cur_privilege r;
- x31 := x31 r;
- x30 := x30 r;
- x29 := x29 r;
- x28 := x28 r;
- x27 := x27 r;
- x26 := x26 r;
- x25 := x25 r;
- x24 := x24 r;
- x23 := x23 r;
- x22 := x22 r;
- x21 := x21 r;
- x20 := x20 r;
- x19 := x19 r;
- x18 := x18 r;
- x17 := x17 r;
- x16 := x16 r;
- x15 := x15 r;
- x14 := x14 r;
- x13 := x13 r;
- x12 := x12 r;
- x11 := x11 r;
- x10 := x10 r;
- x9 := x9 r;
- x8 := x8 r;
- x7 := x7 r;
- x6 := x6 r;
- x5 := x5 r;
- x4 := x4 r;
- x3 := x3 r;
- x2 := x2 r;
- x1 := x1 r;
- Xs := Xs r;
- instbits := instbits r;
- nextPC := nextPC r
- |}.
-
-
-
-Definition Counteren_of_regval (merge_var : register_value)
-: option Counteren :=
-
- match merge_var with | Regval_Counteren (v) => Some (v) | _ => None end.
-
-Definition regval_of_Counteren (v : Counteren) : register_value := Regval_Counteren (v).
-
-Definition Mcause_of_regval (merge_var : register_value)
-: option Mcause :=
-
- match merge_var with | Regval_Mcause (v) => Some (v) | _ => None end.
-
-Definition regval_of_Mcause (v : Mcause) : register_value := Regval_Mcause (v).
-
-Definition Medeleg_of_regval (merge_var : register_value)
-: option Medeleg :=
-
- match merge_var with | Regval_Medeleg (v) => Some (v) | _ => None end.
-
-Definition regval_of_Medeleg (v : Medeleg) : register_value := Regval_Medeleg (v).
-
-Definition Minterrupts_of_regval (merge_var : register_value)
-: option Minterrupts :=
-
- match merge_var with | Regval_Minterrupts (v) => Some (v) | _ => None end.
-
-Definition regval_of_Minterrupts (v : Minterrupts) : register_value := Regval_Minterrupts (v).
-
-Definition Misa_of_regval (merge_var : register_value)
-: option Misa :=
-
- match merge_var with | Regval_Misa (v) => Some (v) | _ => None end.
-
-Definition regval_of_Misa (v : Misa) : register_value := Regval_Misa (v).
-
-Definition Mstatus_of_regval (merge_var : register_value)
-: option Mstatus :=
-
- match merge_var with | Regval_Mstatus (v) => Some (v) | _ => None end.
-
-Definition regval_of_Mstatus (v : Mstatus) : register_value := Regval_Mstatus (v).
-
-Definition Mtvec_of_regval (merge_var : register_value)
-: option Mtvec :=
-
- match merge_var with | Regval_Mtvec (v) => Some (v) | _ => None end.
-
-Definition regval_of_Mtvec (v : Mtvec) : register_value := Regval_Mtvec (v).
-
-Definition Pmpcfg_ent_of_regval (merge_var : register_value)
-: option Pmpcfg_ent :=
-
- match merge_var with | Regval_Pmpcfg_ent (v) => Some (v) | _ => None end.
-
-Definition regval_of_Pmpcfg_ent (v : Pmpcfg_ent) : register_value := Regval_Pmpcfg_ent (v).
-
-Definition Privilege_of_regval (merge_var : register_value)
-: option Privilege :=
-
- match merge_var with | Regval_Privilege (v) => Some (v) | _ => None end.
-
-Definition regval_of_Privilege (v : Privilege) : register_value := Regval_Privilege (v).
-
-Definition Sedeleg_of_regval (merge_var : register_value)
-: option Sedeleg :=
-
- match merge_var with | Regval_Sedeleg (v) => Some (v) | _ => None end.
-
-Definition regval_of_Sedeleg (v : Sedeleg) : register_value := Regval_Sedeleg (v).
-
-Definition Sinterrupts_of_regval (merge_var : register_value)
-: option Sinterrupts :=
-
- match merge_var with | Regval_Sinterrupts (v) => Some (v) | _ => None end.
-
-Definition regval_of_Sinterrupts (v : Sinterrupts) : register_value := Regval_Sinterrupts (v).
+Notation "{[ r 'with' 'satp' := e ]}" :=
+ match r with Build_regstate _ f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate e f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'tlb48' := e ]}" :=
+ match r with Build_regstate f0 _ f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 e f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'tlb39' := e ]}" :=
+ match r with Build_regstate f0 f1 _ f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 e f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'htif_exit_code' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 _ f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 e f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'htif_done' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 _ f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 e f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'htif_tohost' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 _ f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 e f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mtimecmp' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 _ f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 e f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'fcsr' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 _ f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 e f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f31' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 _ f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 e f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f30' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 _ f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 e f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f29' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 _ f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 e f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f28' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 _ f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 e f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f27' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 _ f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 e f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f26' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 _ f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 e f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f25' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 _ f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 e f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f24' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 _ f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 e f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f23' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 _ f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 e f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f22' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 _ f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 e f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f21' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 _ f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 e f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f20' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 _ f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 e f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f19' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 _ f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 e f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f18' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 _ f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 e f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f17' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 _ f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 e f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f16' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 _ f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 e f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f15' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 _ f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 e f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f14' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 _ f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 e f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f13' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 _ f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 e f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f12' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 _ f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 e f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f11' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 _ f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 e f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f10' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 _ f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 e f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f9' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 _ f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 e f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f8' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 _ f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 e f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f7' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 _ f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 e f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f6' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 _ f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 e f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f5' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 _ f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 e f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f4' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 _ f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 e f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f3' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 _ f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 e f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f2' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 _ f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 e f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f1' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 _ f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 e f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'f0' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 _ f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 e f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'fs' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 _ f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 e f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'float_fflags' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 _ f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 e f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'float_result' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 _ f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 e f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'utval' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 _ f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 e f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'ucause' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 _ f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 e f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'uepc' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 _ f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 e f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'uscratch' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 _ f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 e f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'utvec' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 _ f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 e f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr15' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 _ f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 e f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr14' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 _ f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 e f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr13' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 _ f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 e f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr12' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 _ f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 e f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr11' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 _ f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 e f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr10' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 _ f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 e f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr9' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 _ f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 e f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr8' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 _ f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 e f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr7' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 _ f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 e f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr6' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 _ f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 e f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr5' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 _ f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 e f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr4' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 _ f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 e f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr3' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 _ f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 e f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr2' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 _ f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 e f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr1' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 _ f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 e f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmpaddr0' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 _ f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 e f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp15cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 _ f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 e f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp14cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 _ f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 e f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp13cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 _ f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 e f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp12cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 _ f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 e f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp11cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 _ f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 e f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp10cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 _ f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 e f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp9cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 _ f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 e f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp8cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 _ f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 e f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp7cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 _ f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 e f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp6cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 _ f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 e f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp5cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 _ f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 e f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp4cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 _ f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 e f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp3cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 _ f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 e f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp2cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 _ f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 e f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp1cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 _ f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 e f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'pmp0cfg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 _ f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 e f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'tselect' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 _ f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 e f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'stval' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 _ f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 e f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'scause' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 _ f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 e f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'sepc' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 _ f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 e f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'sscratch' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 _ f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 e f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'stvec' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 _ f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 e f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'sideleg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 _ f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 e f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'sedeleg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 _ f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 e f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mhartid' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 _ f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 e f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'marchid' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 _ f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 e f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mimpid' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 _ f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 e f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mvendorid' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 _ f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 e f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'minstret_written' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 _ f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 e f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'minstret' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 _ f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 e f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mtime' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 _ f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 e f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mcycle' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 _ f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 e f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mcountinhibit' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 _ f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 e f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'scounteren' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 _ f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 e f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mcounteren' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 _ f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 e f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mscratch' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 _ f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 e f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mtval' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 _ f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 e f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mepc' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 _ f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 e f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mcause' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 _ f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 e f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mtvec' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 _ f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 e f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'medeleg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 _ f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 e f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mideleg' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 _ f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 e f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mie' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 _ f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 e f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mip' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 _ f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 e f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'mstatus' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 _ f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 e f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'misa' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 _ f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 e f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'cur_inst' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 _ f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 e f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'cur_privilege' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 _ f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 e f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x31' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 _ f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 e f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x30' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 _ f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 e f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x29' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 _ f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 e f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x28' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 _ f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 e f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x27' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 _ f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 e f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x26' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 _ f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 e f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x25' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 _ f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 e f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x24' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 _ f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 e f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x23' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 _ f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 e f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x22' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 _ f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 e f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x21' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 _ f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 e f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x20' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 _ f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 e f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x19' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 _ f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 e f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x18' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 _ f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 e f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x17' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 _ f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 e f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x16' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 _ f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 e f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x15' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 _ f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 e f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x14' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 _ f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 e f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x13' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 _ f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 e f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x12' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 _ f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 e f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x11' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 _ f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 e f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x10' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 _ f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 e f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x9' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 _ f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 e f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x8' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 _ f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 e f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x7' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 _ f137 f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 e f137 f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x6' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 _ f138 f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 e f138 f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x5' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 _ f139 f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 e f139 f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x4' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 _ f140 f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 e f140 f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x3' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 _ f141 f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 e f141 f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x2' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 _ f142 f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 e f142 f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'x1' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 _ f143 f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 e f143 f144 f145 f146
+ end.
+Notation "{[ r 'with' 'Xs' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 _ f144 f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 e f144 f145 f146
+ end.
+Notation "{[ r 'with' 'instbits' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 _ f145 f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 e f145 f146
+ end.
+Notation "{[ r 'with' 'nextPC' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 _ f146 =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 e f146
+ end.
+Notation "{[ r 'with' 'PC' := e ]}" :=
+ match r with Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 _ =>
+ Build_regstate f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 f19 f20 f21 f22 f23 f24 f25 f26 f27 f28 f29 f30 f31 f32 f33 f34 f35 f36 f37 f38 f39 f40 f41 f42 f43 f44 f45 f46 f47 f48 f49 f50 f51 f52 f53 f54 f55 f56 f57 f58 f59 f60 f61 f62 f63 f64 f65 f66 f67 f68 f69 f70 f71 f72 f73 f74 f75 f76 f77 f78 f79 f80 f81 f82 f83 f84 f85 f86 f87 f88 f89 f90 f91 f92 f93 f94 f95 f96 f97 f98 f99 f100 f101 f102 f103 f104 f105 f106 f107 f108 f109 f110 f111 f112 f113 f114 f115 f116 f117 f118 f119 f120 f121 f122 f123 f124 f125 f126 f127 f128 f129 f130 f131 f132 f133 f134 f135 f136 f137 f138 f139 f140 f141 f142 f143 f144 f145 e
+ end.
+
+
+
+Definition Counteren_of_regval (merge_var : register_value) : option Counteren :=
+ match merge_var with | Regval_Counteren v => Some v | _ => None end.
+
+Definition regval_of_Counteren (v : Counteren) : register_value := Regval_Counteren v.
+
+Definition Counterin_of_regval (merge_var : register_value) : option Counterin :=
+ match merge_var with | Regval_Counterin v => Some v | _ => None end.
+
+Definition regval_of_Counterin (v : Counterin) : register_value := Regval_Counterin v.
+
+Definition Fcsr_of_regval (merge_var : register_value) : option Fcsr :=
+ match merge_var with | Regval_Fcsr v => Some v | _ => None end.
+
+Definition regval_of_Fcsr (v : Fcsr) : register_value := Regval_Fcsr v.
+
+Definition Mcause_of_regval (merge_var : register_value) : option Mcause :=
+ match merge_var with | Regval_Mcause v => Some v | _ => None end.
+
+Definition regval_of_Mcause (v : Mcause) : register_value := Regval_Mcause v.
+
+Definition Medeleg_of_regval (merge_var : register_value) : option Medeleg :=
+ match merge_var with | Regval_Medeleg v => Some v | _ => None end.
+
+Definition regval_of_Medeleg (v : Medeleg) : register_value := Regval_Medeleg v.
+
+Definition Minterrupts_of_regval (merge_var : register_value) : option Minterrupts :=
+ match merge_var with | Regval_Minterrupts v => Some v | _ => None end.
+
+Definition regval_of_Minterrupts (v : Minterrupts) : register_value := Regval_Minterrupts v.
+
+Definition Misa_of_regval (merge_var : register_value) : option Misa :=
+ match merge_var with | Regval_Misa v => Some v | _ => None end.
+
+Definition regval_of_Misa (v : Misa) : register_value := Regval_Misa v.
+
+Definition Mstatus_of_regval (merge_var : register_value) : option Mstatus :=
+ match merge_var with | Regval_Mstatus v => Some v | _ => None end.
+
+Definition regval_of_Mstatus (v : Mstatus) : register_value := Regval_Mstatus v.
+
+Definition Mtvec_of_regval (merge_var : register_value) : option Mtvec :=
+ match merge_var with | Regval_Mtvec v => Some v | _ => None end.
+
+Definition regval_of_Mtvec (v : Mtvec) : register_value := Regval_Mtvec v.
+
+Definition Pmpcfg_ent_of_regval (merge_var : register_value) : option Pmpcfg_ent :=
+ match merge_var with | Regval_Pmpcfg_ent v => Some v | _ => None end.
-Definition TLB_Entry_16_39_56_64_of_regval (merge_var : register_value)
+Definition regval_of_Pmpcfg_ent (v : Pmpcfg_ent) : register_value := Regval_Pmpcfg_ent v.
+
+Definition Privilege_of_regval (merge_var : register_value) : option Privilege :=
+ match merge_var with | Regval_Privilege v => Some v | _ => None end.
+
+Definition regval_of_Privilege (v : Privilege) : register_value := Regval_Privilege v.
+
+Definition Sedeleg_of_regval (merge_var : register_value) : option Sedeleg :=
+ match merge_var with | Regval_Sedeleg v => Some v | _ => None end.
+
+Definition regval_of_Sedeleg (v : Sedeleg) : register_value := Regval_Sedeleg v.
+
+Definition Sinterrupts_of_regval (merge_var : register_value) : option Sinterrupts :=
+ match merge_var with | Regval_Sinterrupts v => Some v | _ => None end.
+
+Definition regval_of_Sinterrupts (v : Sinterrupts) : register_value := Regval_Sinterrupts v.
+
+Definition TLB_Entry_16_39_56_64_of_regval (merge_var : register_value)
: option (TLB_Entry 16 39 56 64) :=
-
- match merge_var with | Regval_TLB_Entry_16_39_56_64 (v) => Some (v) | _ => None end.
+ match merge_var with | Regval_TLB_Entry_16_39_56_64 v => Some v | _ => None end.
-Definition regval_of_TLB_Entry_16_39_56_64 (v : TLB_Entry 16 39 56 64)
-: register_value :=
-
- Regval_TLB_Entry_16_39_56_64
- (v).
+Definition regval_of_TLB_Entry_16_39_56_64 (v : TLB_Entry 16 39 56 64) : register_value :=
+ Regval_TLB_Entry_16_39_56_64 v.
-Definition TLB_Entry_16_48_56_64_of_regval (merge_var : register_value)
+Definition TLB_Entry_16_48_56_64_of_regval (merge_var : register_value)
: option (TLB_Entry 16 48 56 64) :=
-
- match merge_var with | Regval_TLB_Entry_16_48_56_64 (v) => Some (v) | _ => None end.
+ match merge_var with | Regval_TLB_Entry_16_48_56_64 v => Some v | _ => None end.
-Definition regval_of_TLB_Entry_16_48_56_64 (v : TLB_Entry 16 48 56 64)
-: register_value :=
-
- Regval_TLB_Entry_16_48_56_64
- (v).
+Definition regval_of_TLB_Entry_16_48_56_64 (v : TLB_Entry 16 48 56 64) : register_value :=
+ Regval_TLB_Entry_16_48_56_64 v.
-Definition bit_of_regval (merge_var : register_value)
-: option bitU :=
-
- match merge_var with | Regval_bit (v) => Some (v) | _ => None end.
+Definition bit_of_regval (merge_var : register_value) : option bitU :=
+ match merge_var with | Regval_bit v => Some v | _ => None end.
-Definition regval_of_bit (v : bitU) : register_value := Regval_bit (v).
+Definition regval_of_bit (v : bitU) : register_value := Regval_bit v.
-Definition bitvector_32_dec_of_regval (merge_var : register_value)
-: option (mword 32) :=
-
- match merge_var with | Regval_bitvector_32_dec (v) => Some (v) | _ => None end.
+Definition bitvector_32_dec_of_regval (merge_var : register_value) : option (mword 32) :=
+ match merge_var with | Regval_bitvector_32_dec v => Some v | _ => None end.
-Definition regval_of_bitvector_32_dec (v : mword 32)
-: register_value :=
-
- Regval_bitvector_32_dec
- (v).
+Definition regval_of_bitvector_32_dec (v : mword 32) : register_value := Regval_bitvector_32_dec v.
-Definition bitvector_64_dec_of_regval (merge_var : register_value)
-: option (mword 64) :=
-
- match merge_var with | Regval_bitvector_64_dec (v) => Some (v) | _ => None end.
+Definition bitvector_64_dec_of_regval (merge_var : register_value) : option (mword 64) :=
+ match merge_var with | Regval_bitvector_64_dec v => Some v | _ => None end.
-Definition regval_of_bitvector_64_dec (v : mword 64)
-: register_value :=
-
- Regval_bitvector_64_dec
- (v).
+Definition regval_of_bitvector_64_dec (v : mword 64) : register_value := Regval_bitvector_64_dec v.
-Definition bool_of_regval (merge_var : register_value)
-: option bool :=
-
- match merge_var with | Regval_bool (v) => Some (v) | _ => None end.
+Definition bool_of_regval (merge_var : register_value) : option bool :=
+ match merge_var with | Regval_bool v => Some v | _ => None end.
-Definition regval_of_bool (v : bool) : register_value := Regval_bool (v).
+Definition regval_of_bool (v : bool) : register_value := Regval_bool v.
Definition vector_of_regval {a} n (of_regval : register_value -> option a) (rv : register_value) : option (vec a n) := match rv with
- | Regval_vector (n', _, v) => if n =? n' then map_bind (vec_of_list n) (just_list (List.map of_regval v)) else None
+ | Regval_vector v => if n =? length_list v then map_bind (vec_of_list n) (just_list (List.map of_regval v)) else None
| _ => None
end.
-Definition regval_of_vector {a} (regval_of : a -> register_value) (size : Z) (is_inc : bool) (xs : vec a size) : register_value := Regval_vector (size, is_inc, List.map regval_of (list_of_vec xs)).
+Definition regval_of_vector {a size} (regval_of : a -> register_value) (xs : vec a size) : register_value := Regval_vector (List.map regval_of (list_of_vec xs)).
Definition list_of_regval {a} (of_regval : register_value -> option a) (rv : register_value) : option (list a) := match rv with
| Regval_list v => just_list (List.map of_regval v)
@@ -13363,6 +1712,258 @@ Definition mtimecmp_ref := {|
of_regval := (fun v => bitvector_64_dec_of_regval v);
regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+Definition fcsr_ref := {|
+ name := "fcsr";
+ read_from := (fun s => s.(fcsr));
+ write_to := (fun v s => ({[ s with fcsr := v ]}));
+ of_regval := (fun v => Fcsr_of_regval v);
+ regval_of := (fun v => regval_of_Fcsr v) |}.
+
+Definition f31_ref := {|
+ name := "f31";
+ read_from := (fun s => s.(f31));
+ write_to := (fun v s => ({[ s with f31 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f30_ref := {|
+ name := "f30";
+ read_from := (fun s => s.(f30));
+ write_to := (fun v s => ({[ s with f30 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f29_ref := {|
+ name := "f29";
+ read_from := (fun s => s.(f29));
+ write_to := (fun v s => ({[ s with f29 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f28_ref := {|
+ name := "f28";
+ read_from := (fun s => s.(f28));
+ write_to := (fun v s => ({[ s with f28 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f27_ref := {|
+ name := "f27";
+ read_from := (fun s => s.(f27));
+ write_to := (fun v s => ({[ s with f27 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f26_ref := {|
+ name := "f26";
+ read_from := (fun s => s.(f26));
+ write_to := (fun v s => ({[ s with f26 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f25_ref := {|
+ name := "f25";
+ read_from := (fun s => s.(f25));
+ write_to := (fun v s => ({[ s with f25 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f24_ref := {|
+ name := "f24";
+ read_from := (fun s => s.(f24));
+ write_to := (fun v s => ({[ s with f24 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f23_ref := {|
+ name := "f23";
+ read_from := (fun s => s.(f23));
+ write_to := (fun v s => ({[ s with f23 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f22_ref := {|
+ name := "f22";
+ read_from := (fun s => s.(f22));
+ write_to := (fun v s => ({[ s with f22 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f21_ref := {|
+ name := "f21";
+ read_from := (fun s => s.(f21));
+ write_to := (fun v s => ({[ s with f21 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f20_ref := {|
+ name := "f20";
+ read_from := (fun s => s.(f20));
+ write_to := (fun v s => ({[ s with f20 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f19_ref := {|
+ name := "f19";
+ read_from := (fun s => s.(f19));
+ write_to := (fun v s => ({[ s with f19 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f18_ref := {|
+ name := "f18";
+ read_from := (fun s => s.(f18));
+ write_to := (fun v s => ({[ s with f18 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f17_ref := {|
+ name := "f17";
+ read_from := (fun s => s.(f17));
+ write_to := (fun v s => ({[ s with f17 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f16_ref := {|
+ name := "f16";
+ read_from := (fun s => s.(f16));
+ write_to := (fun v s => ({[ s with f16 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f15_ref := {|
+ name := "f15";
+ read_from := (fun s => s.(f15));
+ write_to := (fun v s => ({[ s with f15 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f14_ref := {|
+ name := "f14";
+ read_from := (fun s => s.(f14));
+ write_to := (fun v s => ({[ s with f14 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f13_ref := {|
+ name := "f13";
+ read_from := (fun s => s.(f13));
+ write_to := (fun v s => ({[ s with f13 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f12_ref := {|
+ name := "f12";
+ read_from := (fun s => s.(f12));
+ write_to := (fun v s => ({[ s with f12 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f11_ref := {|
+ name := "f11";
+ read_from := (fun s => s.(f11));
+ write_to := (fun v s => ({[ s with f11 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f10_ref := {|
+ name := "f10";
+ read_from := (fun s => s.(f10));
+ write_to := (fun v s => ({[ s with f10 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f9_ref := {|
+ name := "f9";
+ read_from := (fun s => s.(f9));
+ write_to := (fun v s => ({[ s with f9 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f8_ref := {|
+ name := "f8";
+ read_from := (fun s => s.(f8));
+ write_to := (fun v s => ({[ s with f8 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f7_ref := {|
+ name := "f7";
+ read_from := (fun s => s.(f7));
+ write_to := (fun v s => ({[ s with f7 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f6_ref := {|
+ name := "f6";
+ read_from := (fun s => s.(f6));
+ write_to := (fun v s => ({[ s with f6 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f5_ref := {|
+ name := "f5";
+ read_from := (fun s => s.(f5));
+ write_to := (fun v s => ({[ s with f5 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f4_ref := {|
+ name := "f4";
+ read_from := (fun s => s.(f4));
+ write_to := (fun v s => ({[ s with f4 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f3_ref := {|
+ name := "f3";
+ read_from := (fun s => s.(f3));
+ write_to := (fun v s => ({[ s with f3 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f2_ref := {|
+ name := "f2";
+ read_from := (fun s => s.(f2));
+ write_to := (fun v s => ({[ s with f2 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f1_ref := {|
+ name := "f1";
+ read_from := (fun s => s.(f1));
+ write_to := (fun v s => ({[ s with f1 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition f0_ref := {|
+ name := "f0";
+ read_from := (fun s => s.(f0));
+ write_to := (fun v s => ({[ s with f0 := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition fs_ref := {|
+ name := "fs";
+ read_from := (fun s => s.(fs));
+ write_to := (fun v s => ({[ s with fs := v ]}));
+ of_regval := (fun v => vector_of_regval 32 (fun v => bitvector_64_dec_of_regval v) v);
+ regval_of := (fun v => regval_of_vector (fun v => regval_of_bitvector_64_dec v) v) |}.
+
+Definition float_fflags_ref := {|
+ name := "float_fflags";
+ read_from := (fun s => s.(float_fflags));
+ write_to := (fun v s => ({[ s with float_fflags := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
+Definition float_result_ref := {|
+ name := "float_result";
+ read_from := (fun s => s.(float_result));
+ write_to := (fun v s => ({[ s with float_result := v ]}));
+ of_regval := (fun v => bitvector_64_dec_of_regval v);
+ regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+
Definition utval_ref := {|
name := "utval";
read_from := (fun s => s.(utval));
@@ -13734,6 +2335,13 @@ Definition mcycle_ref := {|
of_regval := (fun v => bitvector_64_dec_of_regval v);
regval_of := (fun v => regval_of_bitvector_64_dec v) |}.
+Definition mcountinhibit_ref := {|
+ name := "mcountinhibit";
+ read_from := (fun s => s.(mcountinhibit));
+ write_to := (fun v s => ({[ s with mcountinhibit := v ]}));
+ of_regval := (fun v => Counterin_of_regval v);
+ regval_of := (fun v => regval_of_Counterin v) |}.
+
Definition scounteren_ref := {|
name := "scounteren";
read_from := (fun s => s.(scounteren));
@@ -14061,7 +2669,7 @@ Definition Xs_ref := {|
read_from := (fun s => s.(Xs));
write_to := (fun v s => ({[ s with Xs := v ]}));
of_regval := (fun v => vector_of_regval 32 (fun v => bitvector_64_dec_of_regval v) v);
- regval_of := (fun v => regval_of_vector (fun v => regval_of_bitvector_64_dec v) 32 false v) |}.
+ regval_of := (fun v => regval_of_vector (fun v => regval_of_bitvector_64_dec v) v) |}.
Definition instbits_ref := {|
name := "instbits";
@@ -14093,6 +2701,42 @@ Definition get_regval (reg_name : string) (s : regstate) : option register_value
if string_dec reg_name "htif_done" then Some (htif_done_ref.(regval_of) (htif_done_ref.(read_from) s)) else
if string_dec reg_name "htif_tohost" then Some (htif_tohost_ref.(regval_of) (htif_tohost_ref.(read_from) s)) else
if string_dec reg_name "mtimecmp" then Some (mtimecmp_ref.(regval_of) (mtimecmp_ref.(read_from) s)) else
+ if string_dec reg_name "fcsr" then Some (fcsr_ref.(regval_of) (fcsr_ref.(read_from) s)) else
+ if string_dec reg_name "f31" then Some (f31_ref.(regval_of) (f31_ref.(read_from) s)) else
+ if string_dec reg_name "f30" then Some (f30_ref.(regval_of) (f30_ref.(read_from) s)) else
+ if string_dec reg_name "f29" then Some (f29_ref.(regval_of) (f29_ref.(read_from) s)) else
+ if string_dec reg_name "f28" then Some (f28_ref.(regval_of) (f28_ref.(read_from) s)) else
+ if string_dec reg_name "f27" then Some (f27_ref.(regval_of) (f27_ref.(read_from) s)) else
+ if string_dec reg_name "f26" then Some (f26_ref.(regval_of) (f26_ref.(read_from) s)) else
+ if string_dec reg_name "f25" then Some (f25_ref.(regval_of) (f25_ref.(read_from) s)) else
+ if string_dec reg_name "f24" then Some (f24_ref.(regval_of) (f24_ref.(read_from) s)) else
+ if string_dec reg_name "f23" then Some (f23_ref.(regval_of) (f23_ref.(read_from) s)) else
+ if string_dec reg_name "f22" then Some (f22_ref.(regval_of) (f22_ref.(read_from) s)) else
+ if string_dec reg_name "f21" then Some (f21_ref.(regval_of) (f21_ref.(read_from) s)) else
+ if string_dec reg_name "f20" then Some (f20_ref.(regval_of) (f20_ref.(read_from) s)) else
+ if string_dec reg_name "f19" then Some (f19_ref.(regval_of) (f19_ref.(read_from) s)) else
+ if string_dec reg_name "f18" then Some (f18_ref.(regval_of) (f18_ref.(read_from) s)) else
+ if string_dec reg_name "f17" then Some (f17_ref.(regval_of) (f17_ref.(read_from) s)) else
+ if string_dec reg_name "f16" then Some (f16_ref.(regval_of) (f16_ref.(read_from) s)) else
+ if string_dec reg_name "f15" then Some (f15_ref.(regval_of) (f15_ref.(read_from) s)) else
+ if string_dec reg_name "f14" then Some (f14_ref.(regval_of) (f14_ref.(read_from) s)) else
+ if string_dec reg_name "f13" then Some (f13_ref.(regval_of) (f13_ref.(read_from) s)) else
+ if string_dec reg_name "f12" then Some (f12_ref.(regval_of) (f12_ref.(read_from) s)) else
+ if string_dec reg_name "f11" then Some (f11_ref.(regval_of) (f11_ref.(read_from) s)) else
+ if string_dec reg_name "f10" then Some (f10_ref.(regval_of) (f10_ref.(read_from) s)) else
+ if string_dec reg_name "f9" then Some (f9_ref.(regval_of) (f9_ref.(read_from) s)) else
+ if string_dec reg_name "f8" then Some (f8_ref.(regval_of) (f8_ref.(read_from) s)) else
+ if string_dec reg_name "f7" then Some (f7_ref.(regval_of) (f7_ref.(read_from) s)) else
+ if string_dec reg_name "f6" then Some (f6_ref.(regval_of) (f6_ref.(read_from) s)) else
+ if string_dec reg_name "f5" then Some (f5_ref.(regval_of) (f5_ref.(read_from) s)) else
+ if string_dec reg_name "f4" then Some (f4_ref.(regval_of) (f4_ref.(read_from) s)) else
+ if string_dec reg_name "f3" then Some (f3_ref.(regval_of) (f3_ref.(read_from) s)) else
+ if string_dec reg_name "f2" then Some (f2_ref.(regval_of) (f2_ref.(read_from) s)) else
+ if string_dec reg_name "f1" then Some (f1_ref.(regval_of) (f1_ref.(read_from) s)) else
+ if string_dec reg_name "f0" then Some (f0_ref.(regval_of) (f0_ref.(read_from) s)) else
+ if string_dec reg_name "fs" then Some (fs_ref.(regval_of) (fs_ref.(read_from) s)) else
+ if string_dec reg_name "float_fflags" then Some (float_fflags_ref.(regval_of) (float_fflags_ref.(read_from) s)) else
+ if string_dec reg_name "float_result" then Some (float_result_ref.(regval_of) (float_result_ref.(read_from) s)) else
if string_dec reg_name "utval" then Some (utval_ref.(regval_of) (utval_ref.(read_from) s)) else
if string_dec reg_name "ucause" then Some (ucause_ref.(regval_of) (ucause_ref.(read_from) s)) else
if string_dec reg_name "uepc" then Some (uepc_ref.(regval_of) (uepc_ref.(read_from) s)) else
@@ -14146,6 +2790,7 @@ Definition get_regval (reg_name : string) (s : regstate) : option register_value
if string_dec reg_name "minstret" then Some (minstret_ref.(regval_of) (minstret_ref.(read_from) s)) else
if string_dec reg_name "mtime" then Some (mtime_ref.(regval_of) (mtime_ref.(read_from) s)) else
if string_dec reg_name "mcycle" then Some (mcycle_ref.(regval_of) (mcycle_ref.(read_from) s)) else
+ if string_dec reg_name "mcountinhibit" then Some (mcountinhibit_ref.(regval_of) (mcountinhibit_ref.(read_from) s)) else
if string_dec reg_name "scounteren" then Some (scounteren_ref.(regval_of) (scounteren_ref.(read_from) s)) else
if string_dec reg_name "mcounteren" then Some (mcounteren_ref.(regval_of) (mcounteren_ref.(read_from) s)) else
if string_dec reg_name "mscratch" then Some (mscratch_ref.(regval_of) (mscratch_ref.(read_from) s)) else
@@ -14206,6 +2851,42 @@ Definition set_regval (reg_name : string) (v : register_value) (s : regstate) :
if string_dec reg_name "htif_done" then option_map (fun v => htif_done_ref.(write_to) v s) (htif_done_ref.(of_regval) v) else
if string_dec reg_name "htif_tohost" then option_map (fun v => htif_tohost_ref.(write_to) v s) (htif_tohost_ref.(of_regval) v) else
if string_dec reg_name "mtimecmp" then option_map (fun v => mtimecmp_ref.(write_to) v s) (mtimecmp_ref.(of_regval) v) else
+ if string_dec reg_name "fcsr" then option_map (fun v => fcsr_ref.(write_to) v s) (fcsr_ref.(of_regval) v) else
+ if string_dec reg_name "f31" then option_map (fun v => f31_ref.(write_to) v s) (f31_ref.(of_regval) v) else
+ if string_dec reg_name "f30" then option_map (fun v => f30_ref.(write_to) v s) (f30_ref.(of_regval) v) else
+ if string_dec reg_name "f29" then option_map (fun v => f29_ref.(write_to) v s) (f29_ref.(of_regval) v) else
+ if string_dec reg_name "f28" then option_map (fun v => f28_ref.(write_to) v s) (f28_ref.(of_regval) v) else
+ if string_dec reg_name "f27" then option_map (fun v => f27_ref.(write_to) v s) (f27_ref.(of_regval) v) else
+ if string_dec reg_name "f26" then option_map (fun v => f26_ref.(write_to) v s) (f26_ref.(of_regval) v) else
+ if string_dec reg_name "f25" then option_map (fun v => f25_ref.(write_to) v s) (f25_ref.(of_regval) v) else
+ if string_dec reg_name "f24" then option_map (fun v => f24_ref.(write_to) v s) (f24_ref.(of_regval) v) else
+ if string_dec reg_name "f23" then option_map (fun v => f23_ref.(write_to) v s) (f23_ref.(of_regval) v) else
+ if string_dec reg_name "f22" then option_map (fun v => f22_ref.(write_to) v s) (f22_ref.(of_regval) v) else
+ if string_dec reg_name "f21" then option_map (fun v => f21_ref.(write_to) v s) (f21_ref.(of_regval) v) else
+ if string_dec reg_name "f20" then option_map (fun v => f20_ref.(write_to) v s) (f20_ref.(of_regval) v) else
+ if string_dec reg_name "f19" then option_map (fun v => f19_ref.(write_to) v s) (f19_ref.(of_regval) v) else
+ if string_dec reg_name "f18" then option_map (fun v => f18_ref.(write_to) v s) (f18_ref.(of_regval) v) else
+ if string_dec reg_name "f17" then option_map (fun v => f17_ref.(write_to) v s) (f17_ref.(of_regval) v) else
+ if string_dec reg_name "f16" then option_map (fun v => f16_ref.(write_to) v s) (f16_ref.(of_regval) v) else
+ if string_dec reg_name "f15" then option_map (fun v => f15_ref.(write_to) v s) (f15_ref.(of_regval) v) else
+ if string_dec reg_name "f14" then option_map (fun v => f14_ref.(write_to) v s) (f14_ref.(of_regval) v) else
+ if string_dec reg_name "f13" then option_map (fun v => f13_ref.(write_to) v s) (f13_ref.(of_regval) v) else
+ if string_dec reg_name "f12" then option_map (fun v => f12_ref.(write_to) v s) (f12_ref.(of_regval) v) else
+ if string_dec reg_name "f11" then option_map (fun v => f11_ref.(write_to) v s) (f11_ref.(of_regval) v) else
+ if string_dec reg_name "f10" then option_map (fun v => f10_ref.(write_to) v s) (f10_ref.(of_regval) v) else
+ if string_dec reg_name "f9" then option_map (fun v => f9_ref.(write_to) v s) (f9_ref.(of_regval) v) else
+ if string_dec reg_name "f8" then option_map (fun v => f8_ref.(write_to) v s) (f8_ref.(of_regval) v) else
+ if string_dec reg_name "f7" then option_map (fun v => f7_ref.(write_to) v s) (f7_ref.(of_regval) v) else
+ if string_dec reg_name "f6" then option_map (fun v => f6_ref.(write_to) v s) (f6_ref.(of_regval) v) else
+ if string_dec reg_name "f5" then option_map (fun v => f5_ref.(write_to) v s) (f5_ref.(of_regval) v) else
+ if string_dec reg_name "f4" then option_map (fun v => f4_ref.(write_to) v s) (f4_ref.(of_regval) v) else
+ if string_dec reg_name "f3" then option_map (fun v => f3_ref.(write_to) v s) (f3_ref.(of_regval) v) else
+ if string_dec reg_name "f2" then option_map (fun v => f2_ref.(write_to) v s) (f2_ref.(of_regval) v) else
+ if string_dec reg_name "f1" then option_map (fun v => f1_ref.(write_to) v s) (f1_ref.(of_regval) v) else
+ if string_dec reg_name "f0" then option_map (fun v => f0_ref.(write_to) v s) (f0_ref.(of_regval) v) else
+ if string_dec reg_name "fs" then option_map (fun v => fs_ref.(write_to) v s) (fs_ref.(of_regval) v) else
+ if string_dec reg_name "float_fflags" then option_map (fun v => float_fflags_ref.(write_to) v s) (float_fflags_ref.(of_regval) v) else
+ if string_dec reg_name "float_result" then option_map (fun v => float_result_ref.(write_to) v s) (float_result_ref.(of_regval) v) else
if string_dec reg_name "utval" then option_map (fun v => utval_ref.(write_to) v s) (utval_ref.(of_regval) v) else
if string_dec reg_name "ucause" then option_map (fun v => ucause_ref.(write_to) v s) (ucause_ref.(of_regval) v) else
if string_dec reg_name "uepc" then option_map (fun v => uepc_ref.(write_to) v s) (uepc_ref.(of_regval) v) else
@@ -14259,6 +2940,7 @@ Definition set_regval (reg_name : string) (v : register_value) (s : regstate) :
if string_dec reg_name "minstret" then option_map (fun v => minstret_ref.(write_to) v s) (minstret_ref.(of_regval) v) else
if string_dec reg_name "mtime" then option_map (fun v => mtime_ref.(write_to) v s) (mtime_ref.(of_regval) v) else
if string_dec reg_name "mcycle" then option_map (fun v => mcycle_ref.(write_to) v s) (mcycle_ref.(of_regval) v) else
+ if string_dec reg_name "mcountinhibit" then option_map (fun v => mcountinhibit_ref.(write_to) v s) (mcountinhibit_ref.(of_regval) v) else
if string_dec reg_name "scounteren" then option_map (fun v => scounteren_ref.(write_to) v s) (scounteren_ref.(of_regval) v) else
if string_dec reg_name "mcounteren" then option_map (fun v => mcounteren_ref.(write_to) v s) (mcounteren_ref.(of_regval) v) else
if string_dec reg_name "mscratch" then option_map (fun v => mscratch_ref.(write_to) v s) (mscratch_ref.(of_regval) v) else