chapter \Generated by Lem from \generated_definitions/lem/RV32/riscv.lem\.\ theory "Riscv" imports Main "LEM.Lem_pervasives_extra" "Sail.Sail2_instr_kinds" "Sail.Sail2_values" "Sail.Sail2_operators_mwords" "Sail.Sail2_prompt_monad" "Sail.Sail2_prompt" "Sail.Sail2_string" "Riscv_types" "Mem_metadata" "Riscv_extras_fdext" "Riscv_extras" begin \ \\Generated by Sail from riscv.\\ \ \\open import Pervasives_extra\\ \ \\open import Sail2_instr_kinds\\ \ \\open import Sail2_values\\ \ \\open import Sail2_string\\ \ \\open import Sail2_operators_mwords\\ \ \\open import Sail2_prompt_monad\\ \ \\open import Sail2_prompt\\ \ \\open import Riscv_types\\ \ \\open import Mem_metadata\\ \ \\open import Riscv_extras_fdext\\ \ \\open import Riscv_extras\\ \ \\val is_none : forall 'a. maybe 'a -> bool\\ fun is_none :: \ 'a option \ bool \ where \ is_none (Some (_)) = ( False )\ |\ is_none None = ( True )\ \ \\val is_some : forall 'a. maybe 'a -> bool\\ fun is_some :: \ 'a option \ bool \ where \ is_some (Some (_)) = ( True )\ |\ is_some None = ( False )\ \ \\val eq_unit : unit -> unit -> bool\\ definition eq_unit :: \ unit \ unit \ bool \ where \ eq_unit _ _ = ( True )\ \ \\val neq_bool : bool -> bool -> bool\\ definition neq_bool :: \ bool \ bool \ bool \ where \ neq_bool x y = ( \ (((x = y))))\ for x :: " bool " and y :: " bool " \ \\val __id : integer -> integer\\ definition id0 :: \ int \ int \ where \ id0 x = ( x )\ for x :: " int " \ \\val _shl_int_general : ii -> ii -> ii\\ definition shl_int_general :: \ int \ int \ int \ where \ shl_int_general (m :: ii) (n :: ii) = ( if ((n \ (( 0 :: int)::ii))) then shl_int m n else shr_int m ((- n)))\ for m :: " int " and n :: " int " \ \\val _shr_int_general : ii -> ii -> ii\\ definition shr_int_general :: \ int \ int \ int \ where \ shr_int_general (m :: ii) (n :: ii) = ( if ((n \ (( 0 :: int)::ii))) then shr_int m n else shl_int m ((- n)))\ for m :: " int " and n :: " int " \ \\val fdiv_int : ii -> ii -> ii\\ definition fdiv_int :: \ int \ int \ int \ where \ fdiv_int (n :: ii) (m :: ii) = ( if (((((n < (( 0 :: int)::ii))) \ ((m > (( 0 :: int)::ii)))))) then ((tdiv_int ((n + (( 1 :: int)::ii))) m)) - (( 1 :: int)::ii) else if (((((n > (( 0 :: int)::ii))) \ ((m < (( 0 :: int)::ii)))))) then ((tdiv_int ((n - (( 1 :: int)::ii))) m)) - (( 1 :: int)::ii) else tdiv_int n m )\ for n :: " int " and m :: " int " \ \\val fmod_int : ii -> ii -> ii\\ definition fmod_int :: \ int \ int \ int \ where \ fmod_int (n :: ii) (m :: ii) = ( n - ((m * ((fdiv_int n m)))))\ for n :: " int " and m :: " int " \ \\val concat_str_bits : forall 'n. Size 'n => string -> mword 'n -> string\\ definition concat_str_bits :: \ string \('n::len)Word.word \ string \ where \ concat_str_bits str x = ( (@) str ((string_of_bits x)))\ for str :: " string " and x :: "('n::len)Word.word " \ \\val concat_str_dec : string -> ii -> string\\ definition concat_str_dec :: \ string \ int \ string \ where \ concat_str_dec str x = ( (@) str ((dec_str x)))\ for str :: " string " and x :: " int " \ \\val sail_mask : forall 'len 'v. Size 'len, Size 'v => integer -> mword 'v -> mword 'len\\ definition sail_mask :: \ int \('v::len)Word.word \('len::len)Word.word \ where \ sail_mask len v = ( if ((len \ ((int (size v))))) then (vector_truncate v len :: ( 'len::len)Word.word) else (zero_extend v len :: ( 'len::len)Word.word))\ for len :: " int " and v :: "('v::len)Word.word " \ \\val sail_ones : forall 'n. Size 'n => integer -> mword 'n\\ definition sail_ones :: \ int \('n::len)Word.word \ where \ sail_ones n = ( (not_vec ((zeros n :: ( 'n::len)Word.word)) :: ( 'n::len)Word.word))\ for n :: " int " \ \\val slice_mask : forall 'n. Size 'n => integer -> ii -> ii -> mword 'n\\ definition slice_mask :: \ int \ int \ int \('n::len)Word.word \ where \ slice_mask n i l = ( if ((l \ n)) then (shiftl ((sail_ones n :: ( 'n::len)Word.word)) i :: ( 'n::len)Word.word) else (let (one :: ( 'n::len)Word.word) = ((sail_mask n ( 0b1 :: 1 Word.word) :: ( 'n::len)Word.word)) in (shiftl ((sub_vec ((shiftl one l :: ( 'n::len)Word.word)) one :: ( 'n::len)Word.word)) i :: ( 'n::len)Word.word)))\ for n :: " int " and i :: " int " and l :: " int " \ \\val read_kind_of_num : integer -> read_kind\\ definition read_kind_of_num :: \ int \ read_kind \ where \ read_kind_of_num arg1 = ( (let l__299 = arg1 in if (((l__299 = (( 0 :: int)::ii)))) then Read_plain else if (((l__299 = (( 1 :: int)::ii)))) then Read_reserve else if (((l__299 = (( 2 :: int)::ii)))) then Read_acquire else if (((l__299 = (( 3 :: int)::ii)))) then Read_exclusive else if (((l__299 = (( 4 :: int)::ii)))) then Read_exclusive_acquire else if (((l__299 = (( 5 :: int)::ii)))) then Read_stream else if (((l__299 = (( 6 :: int)::ii)))) then Read_RISCV_acquire else if (((l__299 = (( 7 :: int)::ii)))) then Read_RISCV_strong_acquire else if (((l__299 = (( 8 :: int)::ii)))) then Read_RISCV_reserved else if (((l__299 = (( 9 :: int)::ii)))) then Read_RISCV_reserved_acquire else if (((l__299 = (( 10 :: int)::ii)))) then Read_RISCV_reserved_strong_acquire else Read_X86_locked))\ for arg1 :: " int " \ \\val num_of_read_kind : read_kind -> integer\\ fun num_of_read_kind :: \ read_kind \ int \ where \ num_of_read_kind Read_plain = ( (( 0 :: int)::ii))\ |\ num_of_read_kind Read_reserve = ( (( 1 :: int)::ii))\ |\ num_of_read_kind Read_acquire = ( (( 2 :: int)::ii))\ |\ num_of_read_kind Read_exclusive = ( (( 3 :: int)::ii))\ |\ num_of_read_kind Read_exclusive_acquire = ( (( 4 :: int)::ii))\ |\ num_of_read_kind Read_stream = ( (( 5 :: int)::ii))\ |\ num_of_read_kind Read_RISCV_acquire = ( (( 6 :: int)::ii))\ |\ num_of_read_kind Read_RISCV_strong_acquire = ( (( 7 :: int)::ii))\ |\ num_of_read_kind Read_RISCV_reserved = ( (( 8 :: int)::ii))\ |\ num_of_read_kind Read_RISCV_reserved_acquire = ( (( 9 :: int)::ii))\ |\ num_of_read_kind Read_RISCV_reserved_strong_acquire = ( (( 10 :: int)::ii))\ |\ num_of_read_kind Read_X86_locked = ( (( 11 :: int)::ii))\ \ \\val write_kind_of_num : integer -> write_kind\\ definition write_kind_of_num :: \ int \ write_kind \ where \ write_kind_of_num arg1 = ( (let l__289 = arg1 in if (((l__289 = (( 0 :: int)::ii)))) then Write_plain else if (((l__289 = (( 1 :: int)::ii)))) then Write_conditional else if (((l__289 = (( 2 :: int)::ii)))) then Write_release else if (((l__289 = (( 3 :: int)::ii)))) then Write_exclusive else if (((l__289 = (( 4 :: int)::ii)))) then Write_exclusive_release else if (((l__289 = (( 5 :: int)::ii)))) then Write_RISCV_release else if (((l__289 = (( 6 :: int)::ii)))) then Write_RISCV_strong_release else if (((l__289 = (( 7 :: int)::ii)))) then Write_RISCV_conditional else if (((l__289 = (( 8 :: int)::ii)))) then Write_RISCV_conditional_release else if (((l__289 = (( 9 :: int)::ii)))) then Write_RISCV_conditional_strong_release else Write_X86_locked))\ for arg1 :: " int " \ \\val num_of_write_kind : write_kind -> integer\\ fun num_of_write_kind :: \ write_kind \ int \ where \ num_of_write_kind Write_plain = ( (( 0 :: int)::ii))\ |\ num_of_write_kind Write_conditional = ( (( 1 :: int)::ii))\ |\ num_of_write_kind Write_release = ( (( 2 :: int)::ii))\ |\ num_of_write_kind Write_exclusive = ( (( 3 :: int)::ii))\ |\ num_of_write_kind Write_exclusive_release = ( (( 4 :: int)::ii))\ |\ num_of_write_kind Write_RISCV_release = ( (( 5 :: int)::ii))\ |\ num_of_write_kind Write_RISCV_strong_release = ( (( 6 :: int)::ii))\ |\ num_of_write_kind Write_RISCV_conditional = ( (( 7 :: int)::ii))\ |\ num_of_write_kind Write_RISCV_conditional_release = ( (( 8 :: int)::ii))\ |\ num_of_write_kind Write_RISCV_conditional_strong_release = ( (( 9 :: int)::ii))\ |\ num_of_write_kind Write_X86_locked = ( (( 10 :: int)::ii))\ \ \\val a64_barrier_domain_of_num : integer -> a64_barrier_domain\\ definition a64_barrier_domain_of_num :: \ int \ a64_barrier_domain \ where \ a64_barrier_domain_of_num arg1 = ( (let l__286 = arg1 in if (((l__286 = (( 0 :: int)::ii)))) then A64_FullShare else if (((l__286 = (( 1 :: int)::ii)))) then A64_InnerShare else if (((l__286 = (( 2 :: int)::ii)))) then A64_OuterShare else A64_NonShare))\ for arg1 :: " int " \ \\val num_of_a64_barrier_domain : a64_barrier_domain -> integer\\ fun num_of_a64_barrier_domain :: \ a64_barrier_domain \ int \ where \ num_of_a64_barrier_domain A64_FullShare = ( (( 0 :: int)::ii))\ |\ num_of_a64_barrier_domain A64_InnerShare = ( (( 1 :: int)::ii))\ |\ num_of_a64_barrier_domain A64_OuterShare = ( (( 2 :: int)::ii))\ |\ num_of_a64_barrier_domain A64_NonShare = ( (( 3 :: int)::ii))\ \ \\val a64_barrier_type_of_num : integer -> a64_barrier_type\\ definition a64_barrier_type_of_num :: \ int \ a64_barrier_type \ where \ a64_barrier_type_of_num arg1 = ( (let l__284 = arg1 in if (((l__284 = (( 0 :: int)::ii)))) then A64_barrier_all else if (((l__284 = (( 1 :: int)::ii)))) then A64_barrier_LD else A64_barrier_ST))\ for arg1 :: " int " \ \\val num_of_a64_barrier_type : a64_barrier_type -> integer\\ fun num_of_a64_barrier_type :: \ a64_barrier_type \ int \ where \ num_of_a64_barrier_type A64_barrier_all = ( (( 0 :: int)::ii))\ |\ num_of_a64_barrier_type A64_barrier_LD = ( (( 1 :: int)::ii))\ |\ num_of_a64_barrier_type A64_barrier_ST = ( (( 2 :: int)::ii))\ \ \\val trans_kind_of_num : integer -> trans_kind\\ definition trans_kind_of_num :: \ int \ trans_kind \ where \ trans_kind_of_num arg1 = ( (let l__282 = arg1 in if (((l__282 = (( 0 :: int)::ii)))) then Transaction_start else if (((l__282 = (( 1 :: int)::ii)))) then Transaction_commit else Transaction_abort))\ for arg1 :: " int " \ \\val num_of_trans_kind : trans_kind -> integer\\ fun num_of_trans_kind :: \ trans_kind \ int \ where \ num_of_trans_kind Transaction_start = ( (( 0 :: int)::ii))\ |\ num_of_trans_kind Transaction_commit = ( (( 1 :: int)::ii))\ |\ num_of_trans_kind Transaction_abort = ( (( 2 :: int)::ii))\ \ \\val cache_op_kind_of_num : integer -> cache_op_kind\\ definition cache_op_kind_of_num :: \ int \ cache_op_kind \ where \ cache_op_kind_of_num arg1 = ( (let l__272 = arg1 in if (((l__272 = (( 0 :: int)::ii)))) then Cache_op_D_IVAC else if (((l__272 = (( 1 :: int)::ii)))) then Cache_op_D_ISW else if (((l__272 = (( 2 :: int)::ii)))) then Cache_op_D_CSW else if (((l__272 = (( 3 :: int)::ii)))) then Cache_op_D_CISW else if (((l__272 = (( 4 :: int)::ii)))) then Cache_op_D_ZVA else if (((l__272 = (( 5 :: int)::ii)))) then Cache_op_D_CVAC else if (((l__272 = (( 6 :: int)::ii)))) then Cache_op_D_CVAU else if (((l__272 = (( 7 :: int)::ii)))) then Cache_op_D_CIVAC else if (((l__272 = (( 8 :: int)::ii)))) then Cache_op_I_IALLUIS else if (((l__272 = (( 9 :: int)::ii)))) then Cache_op_I_IALLU else Cache_op_I_IVAU))\ for arg1 :: " int " \ \\val num_of_cache_op_kind : cache_op_kind -> integer\\ fun num_of_cache_op_kind :: \ cache_op_kind \ int \ where \ num_of_cache_op_kind Cache_op_D_IVAC = ( (( 0 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_D_ISW = ( (( 1 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_D_CSW = ( (( 2 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_D_CISW = ( (( 3 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_D_ZVA = ( (( 4 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_D_CVAC = ( (( 5 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_D_CVAU = ( (( 6 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_D_CIVAC = ( (( 7 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_I_IALLUIS = ( (( 8 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_I_IALLU = ( (( 9 :: int)::ii))\ |\ num_of_cache_op_kind Cache_op_I_IVAU = ( (( 10 :: int)::ii))\ \ \\val not_bit : bitU -> bitU\\ definition not_bit0 :: \ bitU \ bitU \ where \ not_bit0 b = ( if (((b = B1))) then B0 else B1 )\ for b :: " bitU " \ \\val string_of_bit : bitU -> string\\ fun string_of_bit :: \ bitU \ string \ where \ string_of_bit B0 = ( (''0b0''))\ |\ string_of_bit B1 = ( (''0b1''))\ \ \\val get_config_print_instr : unit -> bool\\ \ \\val get_config_print_reg : unit -> bool\\ \ \\val get_config_print_mem : unit -> bool\\ \ \\val get_config_print_platform : unit -> bool\\ definition get_config_print_instr :: \ unit \ bool \ where \ get_config_print_instr _ = ( False )\ definition get_config_print_reg :: \ unit \ bool \ where \ get_config_print_reg _ = ( False )\ definition get_config_print_mem :: \ unit \ bool \ where \ get_config_print_mem _ = ( False )\ definition get_config_print_platform :: \ unit \ bool \ where \ get_config_print_platform _ = ( False )\ \ \\val EXTS : forall 'm 'n. Size 'm, Size 'n => integer -> mword 'n -> mword 'm\\ \ \\val EXTZ : forall 'm 'n. Size 'm, Size 'n => integer -> mword 'n -> mword 'm\\ definition EXTS :: \ int \('n::len)Word.word \('m::len)Word.word \ where \ EXTS m v = ( (sign_extend v m :: ( 'm::len)Word.word))\ for m :: " int " and v :: "('n::len)Word.word " definition EXTZ :: \ int \('n::len)Word.word \('m::len)Word.word \ where \ EXTZ m v = ( (zero_extend v m :: ( 'm::len)Word.word))\ for m :: " int " and v :: "('n::len)Word.word " \ \\val zeros_implicit : forall 'n. Size 'n => integer -> mword 'n\\ definition zeros_implicit :: \ int \('n::len)Word.word \ where \ zeros_implicit n = ( (zeros n :: ( 'n::len)Word.word))\ for n :: " int " \ \\val ones : forall 'n. Size 'n => integer -> mword 'n\\ definition ones :: \ int \('n::len)Word.word \ where \ ones n = ( (sail_ones n :: ( 'n::len)Word.word))\ for n :: " int " \ \\val bool_to_bits : bool -> mword ty1\\ definition bool_to_bits :: \ bool \(1)Word.word \ where \ bool_to_bits x = ( if x then ( 0b1 :: 1 Word.word) else ( 0b0 :: 1 Word.word))\ for x :: " bool " \ \\val bit_to_bool : bitU -> bool\\ fun bit_to_bool :: \ bitU \ bool \ where \ bit_to_bool B1 = ( True )\ |\ bit_to_bool B0 = ( False )\ \ \\val to_bits : forall 'l. Size 'l => integer -> ii -> mword 'l\\ definition to_bits :: \ int \ int \('l::len)Word.word \ where \ to_bits l n = ( (get_slice_int l n (( 0 :: int)::ii) :: ( 'l::len)Word.word))\ for l :: " int " and n :: " int " \ \\val zopz0zI_s : forall 'n. Size 'n => mword 'n -> mword 'n -> bool\\ \ \\val zopz0zKzJ_s : forall 'n. Size 'n => mword 'n -> mword 'n -> bool\\ \ \\val zopz0zI_u : forall 'n. Size 'n => mword 'n -> mword 'n -> bool\\ \ \\val zopz0zKzJ_u : forall 'n. Size 'n => mword 'n -> mword 'n -> bool\\ \ \\val zopz0zIzJ_u : forall 'n. Size 'n => mword 'n -> mword 'n -> bool\\ definition zopz0zI_s :: \('n::len)Word.word \('n::len)Word.word \ bool \ where \ zopz0zI_s x y = ( ((Word.sint x)) < ((Word.sint y)))\ for x :: "('n::len)Word.word " and y :: "('n::len)Word.word " definition zopz0zKzJ_s :: \('n::len)Word.word \('n::len)Word.word \ bool \ where \ zopz0zKzJ_s x y = ( ((Word.sint x)) \ ((Word.sint y)))\ for x :: "('n::len)Word.word " and y :: "('n::len)Word.word " definition zopz0zI_u :: \('n::len)Word.word \('n::len)Word.word \ bool \ where \ zopz0zI_u x y = ( ((Word.uint x)) < ((Word.uint y)))\ for x :: "('n::len)Word.word " and y :: "('n::len)Word.word " definition zopz0zKzJ_u :: \('n::len)Word.word \('n::len)Word.word \ bool \ where \ zopz0zKzJ_u x y = ( ((Word.uint x)) \ ((Word.uint y)))\ for x :: "('n::len)Word.word " and y :: "('n::len)Word.word " definition zopz0zIzJ_u :: \('n::len)Word.word \('n::len)Word.word \ bool \ where \ zopz0zIzJ_u x y = ( ((Word.uint x)) \ ((Word.uint y)))\ for x :: "('n::len)Word.word " and y :: "('n::len)Word.word " \ \\val shift_right_arith64 : mword ty64 -> mword ty6 -> mword ty64\\ definition shift_right_arith64 :: \(64)Word.word \(6)Word.word \(64)Word.word \ where \ shift_right_arith64 (v :: 64 bits) (shift :: 6 bits) = ( (let (v128 :: 128 bits) = ((EXTS (( 128 :: int)::ii) v :: 128 Word.word)) in (subrange_vec_dec ((shift_bits_right v128 shift :: 128 Word.word)) (( 63 :: int)::ii) (( 0 :: int)::ii) :: 64 Word.word)))\ for v :: "(64)Word.word " and shift :: "(6)Word.word " \ \\val shift_right_arith32 : mword ty32 -> mword ty5 -> mword ty32\\ definition shift_right_arith32 :: \(32)Word.word \(5)Word.word \(32)Word.word \ where \ shift_right_arith32 (v :: 32 bits) (shift :: 5 bits) = ( (let (v64 :: 64 bits) = ((EXTS (( 64 :: int)::ii) v :: 64 Word.word)) in (subrange_vec_dec ((shift_bits_right v64 shift :: 64 Word.word)) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)))\ for v :: "(32)Word.word " and shift :: "(5)Word.word " \ \\val spc_forwards : unit -> string\\ \ \\val spc_backwards : string -> unit\\ \ \\val spc_forwards_matches : unit -> bool\\ \ \\val spc_backwards_matches : string -> bool\\ \ \\val spc_matches_prefix : string -> maybe ((unit * ii))\\ \ \\val opt_spc_forwards : unit -> string\\ \ \\val opt_spc_backwards : string -> unit\\ \ \\val opt_spc_forwards_matches : unit -> bool\\ \ \\val opt_spc_backwards_matches : string -> bool\\ \ \\val opt_spc_matches_prefix : string -> maybe ((unit * ii))\\ \ \\val def_spc_forwards : unit -> string\\ \ \\val def_spc_backwards : string -> unit\\ \ \\val def_spc_forwards_matches : unit -> bool\\ \ \\val def_spc_backwards_matches : string -> bool\\ \ \\val def_spc_matches_prefix : string -> maybe ((unit * ii))\\ \ \\val hex_bits_forwards : forall 'n. Size 'n => (integer * mword 'n) -> string\\ \ \\val hex_bits_backwards : forall 'n. Size 'n => string -> (integer * mword 'n)\\ \ \\val hex_bits_forwards_matches : forall 'n. Size 'n => (integer * mword 'n) -> bool\\ \ \\val hex_bits_backwards_matches : string -> bool\\ \ \\val hex_bits_matches_prefix : forall 'n. Size 'n => string -> maybe (((integer * mword 'n) * ii))\\ \ \\val n_leading_spaces : string -> ii\\ function (sequential,domintros) n_leading_spaces0 :: \ string \ int \ where \ n_leading_spaces0 s = ( (let p00 = s in if (((p00 = ('''')))) then (( 0 :: int)::ii) else (let p00 = (string_take s (( 1 :: int)::ii)) in if (((p00 = ('' '')))) then (( 1 :: int)::ii) + ((n_leading_spaces0 ((string_drop s (( 1 :: int)::ii))))) else (( 0 :: int)::ii))))\ for s :: " string " by pat_completeness (auto intro!: let_cong bind_cong MemoryOpResult.case_cong) definition spc_forwards :: \ unit \ string \ where \ spc_forwards _ = ( ('' ''))\ definition spc_backwards :: \ string \ unit \ where \ spc_backwards s = ( () )\ for s :: " string " definition spc_matches_prefix0 :: \ string \(unit*int)option \ where \ spc_matches_prefix0 s = ( (let n = (n_leading_spaces0 s) in (let l__271 = n in if (((l__271 = (( 0 :: int)::ii)))) then None else Some (() , n))))\ for s :: " string " definition opt_spc_forwards :: \ unit \ string \ where \ opt_spc_forwards _ = ( (''''))\ definition opt_spc_backwards :: \ string \ unit \ where \ opt_spc_backwards s = ( () )\ for s :: " string " definition opt_spc_matches_prefix0 :: \ string \(unit*int)option \ where \ opt_spc_matches_prefix0 s = ( Some (() , n_leading_spaces0 s))\ for s :: " string " definition def_spc_forwards :: \ unit \ string \ where \ def_spc_forwards _ = ( ('' ''))\ definition def_spc_backwards :: \ string \ unit \ where \ def_spc_backwards s = ( () )\ for s :: " string " definition def_spc_matches_prefix :: \ string \(unit*ii)option \ where \ def_spc_matches_prefix s = ( opt_spc_matches_prefix0 s )\ for s :: " string " \ \\val hex_bits_1_backwards : string -> M (mword ty1)\\ \ \\val hex_bits_1_forwards_matches : mword ty1 -> bool\\ \ \\val hex_bits_1_backwards_matches : string -> bool\\ definition hex_bits_1_forwards_matches :: \(1)Word.word \ bool \ where \ hex_bits_1_forwards_matches bv = ( True )\ for bv :: "(1)Word.word " definition hex_bits_1_backwards_matches :: \ string \ bool \ where \ hex_bits_1_backwards_matches s = ( if ((case ((hex_bits_1_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 1 Word.word * ii))option)) of Some ((g__362, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_1_backwards :: \ string \((register_value),((1)Word.word),(exception))monad \ where \ hex_bits_1_backwards s = ( (case ((hex_bits_1_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 1 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 49:2 - 51:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 49:2 - 51:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_2_backwards : string -> M (mword ty2)\\ \ \\val hex_bits_2_forwards_matches : mword ty2 -> bool\\ \ \\val hex_bits_2_backwards_matches : string -> bool\\ definition hex_bits_2_forwards_matches :: \(2)Word.word \ bool \ where \ hex_bits_2_forwards_matches bv = ( True )\ for bv :: "(2)Word.word " definition hex_bits_2_backwards_matches :: \ string \ bool \ where \ hex_bits_2_backwards_matches s = ( if ((case ((hex_bits_2_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 2 Word.word * ii))option)) of Some ((g__361, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_2_backwards :: \ string \((register_value),((2)Word.word),(exception))monad \ where \ hex_bits_2_backwards s = ( (case ((hex_bits_2_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 2 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 68:2 - 70:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 68:2 - 70:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_3_backwards : string -> M (mword ty3)\\ \ \\val hex_bits_3_forwards_matches : mword ty3 -> bool\\ \ \\val hex_bits_3_backwards_matches : string -> bool\\ definition hex_bits_3_forwards_matches :: \(3)Word.word \ bool \ where \ hex_bits_3_forwards_matches bv = ( True )\ for bv :: "(3)Word.word " definition hex_bits_3_backwards_matches :: \ string \ bool \ where \ hex_bits_3_backwards_matches s = ( if ((case ((hex_bits_3_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 3 Word.word * ii))option)) of Some ((g__360, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_3_backwards :: \ string \((register_value),((3)Word.word),(exception))monad \ where \ hex_bits_3_backwards s = ( (case ((hex_bits_3_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 3 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 87:2 - 89:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 87:2 - 89:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_4_backwards : string -> M (mword ty4)\\ \ \\val hex_bits_4_forwards_matches : mword ty4 -> bool\\ \ \\val hex_bits_4_backwards_matches : string -> bool\\ definition hex_bits_4_forwards_matches :: \(4)Word.word \ bool \ where \ hex_bits_4_forwards_matches bv = ( True )\ for bv :: "(4)Word.word " definition hex_bits_4_backwards_matches :: \ string \ bool \ where \ hex_bits_4_backwards_matches s = ( if ((case ((hex_bits_4_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 4 Word.word * ii))option)) of Some ((g__359, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_4_backwards :: \ string \((register_value),((4)Word.word),(exception))monad \ where \ hex_bits_4_backwards s = ( (case ((hex_bits_4_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 4 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 106:2 - 108:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 106:2 - 108:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_5_backwards : string -> M (mword ty5)\\ \ \\val hex_bits_5_forwards_matches : mword ty5 -> bool\\ \ \\val hex_bits_5_backwards_matches : string -> bool\\ definition hex_bits_5_forwards_matches :: \(5)Word.word \ bool \ where \ hex_bits_5_forwards_matches bv = ( True )\ for bv :: "(5)Word.word " definition hex_bits_5_backwards_matches :: \ string \ bool \ where \ hex_bits_5_backwards_matches s = ( if ((case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 5 Word.word * ii))option)) of Some ((g__358, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_5_backwards :: \ string \((register_value),((5)Word.word),(exception))monad \ where \ hex_bits_5_backwards s = ( (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 5 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 125:2 - 127:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 125:2 - 127:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_6_backwards : string -> M (mword ty6)\\ \ \\val hex_bits_6_forwards_matches : mword ty6 -> bool\\ \ \\val hex_bits_6_backwards_matches : string -> bool\\ definition hex_bits_6_forwards_matches :: \(6)Word.word \ bool \ where \ hex_bits_6_forwards_matches bv = ( True )\ for bv :: "(6)Word.word " definition hex_bits_6_backwards_matches :: \ string \ bool \ where \ hex_bits_6_backwards_matches s = ( if ((case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 6 Word.word * ii))option)) of Some ((g__357, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_6_backwards :: \ string \((register_value),((6)Word.word),(exception))monad \ where \ hex_bits_6_backwards s = ( (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 6 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 144:2 - 146:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 144:2 - 146:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_7_backwards : string -> M (mword ty7)\\ \ \\val hex_bits_7_forwards_matches : mword ty7 -> bool\\ \ \\val hex_bits_7_backwards_matches : string -> bool\\ definition hex_bits_7_forwards_matches :: \(7)Word.word \ bool \ where \ hex_bits_7_forwards_matches bv = ( True )\ for bv :: "(7)Word.word " definition hex_bits_7_backwards_matches :: \ string \ bool \ where \ hex_bits_7_backwards_matches s = ( if ((case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 7 Word.word * ii))option)) of Some ((g__356, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_7_backwards :: \ string \((register_value),((7)Word.word),(exception))monad \ where \ hex_bits_7_backwards s = ( (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 7 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 163:2 - 165:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 163:2 - 165:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_8_backwards : string -> M (mword ty8)\\ \ \\val hex_bits_8_forwards_matches : mword ty8 -> bool\\ \ \\val hex_bits_8_backwards_matches : string -> bool\\ definition hex_bits_8_forwards_matches :: \(8)Word.word \ bool \ where \ hex_bits_8_forwards_matches bv = ( True )\ for bv :: "(8)Word.word " definition hex_bits_8_backwards_matches :: \ string \ bool \ where \ hex_bits_8_backwards_matches s = ( if ((case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 8 Word.word * ii))option)) of Some ((g__355, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_8_backwards :: \ string \((register_value),((8)Word.word),(exception))monad \ where \ hex_bits_8_backwards s = ( (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 8 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 182:2 - 184:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 182:2 - 184:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_9_backwards : string -> M (mword ty9)\\ \ \\val hex_bits_9_forwards_matches : mword ty9 -> bool\\ \ \\val hex_bits_9_backwards_matches : string -> bool\\ definition hex_bits_9_forwards_matches :: \(9)Word.word \ bool \ where \ hex_bits_9_forwards_matches bv = ( True )\ for bv :: "(9)Word.word " definition hex_bits_9_backwards_matches :: \ string \ bool \ where \ hex_bits_9_backwards_matches s = ( if ((case ((hex_bits_9_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 9 Word.word * ii))option)) of Some ((g__354, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_9_backwards :: \ string \((register_value),((9)Word.word),(exception))monad \ where \ hex_bits_9_backwards s = ( (case ((hex_bits_9_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 9 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 201:2 - 203:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 201:2 - 203:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_10_backwards : string -> M (mword ty10)\\ \ \\val hex_bits_10_forwards_matches : mword ty10 -> bool\\ \ \\val hex_bits_10_backwards_matches : string -> bool\\ definition hex_bits_10_forwards_matches :: \(10)Word.word \ bool \ where \ hex_bits_10_forwards_matches bv = ( True )\ for bv :: "(10)Word.word " definition hex_bits_10_backwards_matches :: \ string \ bool \ where \ hex_bits_10_backwards_matches s = ( if ((case ((hex_bits_10_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 10 Word.word * ii))option)) of Some ((g__353, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_10_backwards :: \ string \((register_value),((10)Word.word),(exception))monad \ where \ hex_bits_10_backwards s = ( (case ((hex_bits_10_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 10 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 220:2 - 222:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 220:2 - 222:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_11_backwards : string -> M (mword ty11)\\ \ \\val hex_bits_11_forwards_matches : mword ty11 -> bool\\ \ \\val hex_bits_11_backwards_matches : string -> bool\\ definition hex_bits_11_forwards_matches :: \(11)Word.word \ bool \ where \ hex_bits_11_forwards_matches bv = ( True )\ for bv :: "(11)Word.word " definition hex_bits_11_backwards_matches :: \ string \ bool \ where \ hex_bits_11_backwards_matches s = ( if ((case ((hex_bits_11_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 11 Word.word * ii))option)) of Some ((g__352, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_11_backwards :: \ string \((register_value),((11)Word.word),(exception))monad \ where \ hex_bits_11_backwards s = ( (case ((hex_bits_11_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 11 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 239:2 - 241:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 239:2 - 241:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_12_backwards : string -> M (mword ty12)\\ \ \\val hex_bits_12_forwards_matches : mword ty12 -> bool\\ \ \\val hex_bits_12_backwards_matches : string -> bool\\ \ \\val hex_bits_12_matches_prefix : string -> maybe ((mword ty12 * ii))\\ definition hex_bits_12_forwards_matches :: \(12)Word.word \ bool \ where \ hex_bits_12_forwards_matches bv = ( True )\ for bv :: "(12)Word.word " definition hex_bits_12_matches_prefix0 :: \ string \((12)Word.word*ii)option \ where \ hex_bits_12_matches_prefix0 s = ( None )\ for s :: " string " definition hex_bits_12_backwards_matches :: \ string \ bool \ where \ hex_bits_12_backwards_matches s = ( if ((case ((hex_bits_12_matches_prefix0 s :: (( 12 Word.word * ii))option)) of Some ((g__351, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_12_backwards :: \ string \((register_value),((12)Word.word),(exception))monad \ where \ hex_bits_12_backwards s = ( (case ((hex_bits_12_matches_prefix0 s :: (( 12 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 262:2 - 264:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 262:2 - 264:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_13_backwards : string -> M (mword ty13)\\ \ \\val hex_bits_13_forwards_matches : mword ty13 -> bool\\ \ \\val hex_bits_13_backwards_matches : string -> bool\\ definition hex_bits_13_forwards_matches :: \(13)Word.word \ bool \ where \ hex_bits_13_forwards_matches bv = ( True )\ for bv :: "(13)Word.word " definition hex_bits_13_backwards_matches :: \ string \ bool \ where \ hex_bits_13_backwards_matches s = ( if ((case ((hex_bits_13_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 13 Word.word * ii))option)) of Some ((g__350, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_13_backwards :: \ string \((register_value),((13)Word.word),(exception))monad \ where \ hex_bits_13_backwards s = ( (case ((hex_bits_13_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 13 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 281:2 - 283:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 281:2 - 283:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_14_backwards : string -> M (mword ty14)\\ \ \\val hex_bits_14_forwards_matches : mword ty14 -> bool\\ \ \\val hex_bits_14_backwards_matches : string -> bool\\ definition hex_bits_14_forwards_matches :: \(14)Word.word \ bool \ where \ hex_bits_14_forwards_matches bv = ( True )\ for bv :: "(14)Word.word " definition hex_bits_14_backwards_matches :: \ string \ bool \ where \ hex_bits_14_backwards_matches s = ( if ((case ((hex_bits_14_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 14 Word.word * ii))option)) of Some ((g__349, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_14_backwards :: \ string \((register_value),((14)Word.word),(exception))monad \ where \ hex_bits_14_backwards s = ( (case ((hex_bits_14_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 14 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 300:2 - 302:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 300:2 - 302:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_15_backwards : string -> M (mword ty15)\\ \ \\val hex_bits_15_forwards_matches : mword ty15 -> bool\\ \ \\val hex_bits_15_backwards_matches : string -> bool\\ definition hex_bits_15_forwards_matches :: \(15)Word.word \ bool \ where \ hex_bits_15_forwards_matches bv = ( True )\ for bv :: "(15)Word.word " definition hex_bits_15_backwards_matches :: \ string \ bool \ where \ hex_bits_15_backwards_matches s = ( if ((case ((hex_bits_15_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 15 Word.word * ii))option)) of Some ((g__348, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_15_backwards :: \ string \((register_value),((15)Word.word),(exception))monad \ where \ hex_bits_15_backwards s = ( (case ((hex_bits_15_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 15 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 319:2 - 321:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 319:2 - 321:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_16_backwards : string -> M (mword ty16)\\ \ \\val hex_bits_16_forwards_matches : mword ty16 -> bool\\ \ \\val hex_bits_16_backwards_matches : string -> bool\\ definition hex_bits_16_forwards_matches :: \(16)Word.word \ bool \ where \ hex_bits_16_forwards_matches bv = ( True )\ for bv :: "(16)Word.word " definition hex_bits_16_backwards_matches :: \ string \ bool \ where \ hex_bits_16_backwards_matches s = ( if ((case ((hex_bits_16_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 16 Word.word * ii))option)) of Some ((g__347, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_16_backwards :: \ string \((register_value),((16)Word.word),(exception))monad \ where \ hex_bits_16_backwards s = ( (case ((hex_bits_16_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 16 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 338:2 - 340:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 338:2 - 340:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_17_backwards : string -> M (mword ty17)\\ \ \\val hex_bits_17_forwards_matches : mword ty17 -> bool\\ \ \\val hex_bits_17_backwards_matches : string -> bool\\ definition hex_bits_17_forwards_matches :: \(17)Word.word \ bool \ where \ hex_bits_17_forwards_matches bv = ( True )\ for bv :: "(17)Word.word " definition hex_bits_17_backwards_matches :: \ string \ bool \ where \ hex_bits_17_backwards_matches s = ( if ((case ((hex_bits_17_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 17 Word.word * ii))option)) of Some ((g__346, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_17_backwards :: \ string \((register_value),((17)Word.word),(exception))monad \ where \ hex_bits_17_backwards s = ( (case ((hex_bits_17_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 17 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 357:2 - 359:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 357:2 - 359:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_18_backwards : string -> M (mword ty18)\\ \ \\val hex_bits_18_forwards_matches : mword ty18 -> bool\\ \ \\val hex_bits_18_backwards_matches : string -> bool\\ definition hex_bits_18_forwards_matches :: \(18)Word.word \ bool \ where \ hex_bits_18_forwards_matches bv = ( True )\ for bv :: "(18)Word.word " definition hex_bits_18_backwards_matches :: \ string \ bool \ where \ hex_bits_18_backwards_matches s = ( if ((case ((hex_bits_18_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 18 Word.word * ii))option)) of Some ((g__345, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_18_backwards :: \ string \((register_value),((18)Word.word),(exception))monad \ where \ hex_bits_18_backwards s = ( (case ((hex_bits_18_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 18 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 376:2 - 378:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 376:2 - 378:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_19_backwards : string -> M (mword ty19)\\ \ \\val hex_bits_19_forwards_matches : mword ty19 -> bool\\ \ \\val hex_bits_19_backwards_matches : string -> bool\\ definition hex_bits_19_forwards_matches :: \(19)Word.word \ bool \ where \ hex_bits_19_forwards_matches bv = ( True )\ for bv :: "(19)Word.word " definition hex_bits_19_backwards_matches :: \ string \ bool \ where \ hex_bits_19_backwards_matches s = ( if ((case ((hex_bits_19_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 19 Word.word * ii))option)) of Some ((g__344, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_19_backwards :: \ string \((register_value),((19)Word.word),(exception))monad \ where \ hex_bits_19_backwards s = ( (case ((hex_bits_19_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 19 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 395:2 - 397:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 395:2 - 397:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_20_backwards : string -> M (mword ty20)\\ \ \\val hex_bits_20_forwards_matches : mword ty20 -> bool\\ \ \\val hex_bits_20_backwards_matches : string -> bool\\ definition hex_bits_20_forwards_matches :: \(20)Word.word \ bool \ where \ hex_bits_20_forwards_matches bv = ( True )\ for bv :: "(20)Word.word " definition hex_bits_20_backwards_matches :: \ string \ bool \ where \ hex_bits_20_backwards_matches s = ( if ((case ((hex_bits_20_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 20 Word.word * ii))option)) of Some ((g__343, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_20_backwards :: \ string \((register_value),((20)Word.word),(exception))monad \ where \ hex_bits_20_backwards s = ( (case ((hex_bits_20_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 20 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 414:2 - 416:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 414:2 - 416:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_21_backwards : string -> M (mword ty21)\\ \ \\val hex_bits_21_forwards_matches : mword ty21 -> bool\\ \ \\val hex_bits_21_backwards_matches : string -> bool\\ definition hex_bits_21_forwards_matches :: \(21)Word.word \ bool \ where \ hex_bits_21_forwards_matches bv = ( True )\ for bv :: "(21)Word.word " definition hex_bits_21_backwards_matches :: \ string \ bool \ where \ hex_bits_21_backwards_matches s = ( if ((case ((hex_bits_21_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 21 Word.word * ii))option)) of Some ((g__342, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_21_backwards :: \ string \((register_value),((21)Word.word),(exception))monad \ where \ hex_bits_21_backwards s = ( (case ((hex_bits_21_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 21 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 433:2 - 435:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 433:2 - 435:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_22_backwards : string -> M (mword ty22)\\ \ \\val hex_bits_22_forwards_matches : mword ty22 -> bool\\ \ \\val hex_bits_22_backwards_matches : string -> bool\\ definition hex_bits_22_forwards_matches :: \(22)Word.word \ bool \ where \ hex_bits_22_forwards_matches bv = ( True )\ for bv :: "(22)Word.word " definition hex_bits_22_backwards_matches :: \ string \ bool \ where \ hex_bits_22_backwards_matches s = ( if ((case ((hex_bits_22_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 22 Word.word * ii))option)) of Some ((g__341, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_22_backwards :: \ string \((register_value),((22)Word.word),(exception))monad \ where \ hex_bits_22_backwards s = ( (case ((hex_bits_22_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 22 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 452:2 - 454:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 452:2 - 454:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_23_backwards : string -> M (mword ty23)\\ \ \\val hex_bits_23_forwards_matches : mword ty23 -> bool\\ \ \\val hex_bits_23_backwards_matches : string -> bool\\ definition hex_bits_23_forwards_matches :: \(23)Word.word \ bool \ where \ hex_bits_23_forwards_matches bv = ( True )\ for bv :: "(23)Word.word " definition hex_bits_23_backwards_matches :: \ string \ bool \ where \ hex_bits_23_backwards_matches s = ( if ((case ((hex_bits_23_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 23 Word.word * ii))option)) of Some ((g__340, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_23_backwards :: \ string \((register_value),((23)Word.word),(exception))monad \ where \ hex_bits_23_backwards s = ( (case ((hex_bits_23_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 23 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 471:2 - 473:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 471:2 - 473:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_24_backwards : string -> M (mword ty24)\\ \ \\val hex_bits_24_forwards_matches : mword ty24 -> bool\\ \ \\val hex_bits_24_backwards_matches : string -> bool\\ definition hex_bits_24_forwards_matches :: \(24)Word.word \ bool \ where \ hex_bits_24_forwards_matches bv = ( True )\ for bv :: "(24)Word.word " definition hex_bits_24_backwards_matches :: \ string \ bool \ where \ hex_bits_24_backwards_matches s = ( if ((case ((hex_bits_24_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 24 Word.word * ii))option)) of Some ((g__339, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_24_backwards :: \ string \((register_value),((24)Word.word),(exception))monad \ where \ hex_bits_24_backwards s = ( (case ((hex_bits_24_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 24 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 490:2 - 492:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 490:2 - 492:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_25_backwards : string -> M (mword ty25)\\ \ \\val hex_bits_25_forwards_matches : mword ty25 -> bool\\ \ \\val hex_bits_25_backwards_matches : string -> bool\\ definition hex_bits_25_forwards_matches :: \(25)Word.word \ bool \ where \ hex_bits_25_forwards_matches bv = ( True )\ for bv :: "(25)Word.word " definition hex_bits_25_backwards_matches :: \ string \ bool \ where \ hex_bits_25_backwards_matches s = ( if ((case ((hex_bits_25_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 25 Word.word * ii))option)) of Some ((g__338, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_25_backwards :: \ string \((register_value),((25)Word.word),(exception))monad \ where \ hex_bits_25_backwards s = ( (case ((hex_bits_25_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 25 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 509:2 - 511:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 509:2 - 511:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_26_backwards : string -> M (mword ty26)\\ \ \\val hex_bits_26_forwards_matches : mword ty26 -> bool\\ \ \\val hex_bits_26_backwards_matches : string -> bool\\ definition hex_bits_26_forwards_matches :: \(26)Word.word \ bool \ where \ hex_bits_26_forwards_matches bv = ( True )\ for bv :: "(26)Word.word " definition hex_bits_26_backwards_matches :: \ string \ bool \ where \ hex_bits_26_backwards_matches s = ( if ((case ((hex_bits_26_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 26 Word.word * ii))option)) of Some ((g__337, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_26_backwards :: \ string \((register_value),((26)Word.word),(exception))monad \ where \ hex_bits_26_backwards s = ( (case ((hex_bits_26_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 26 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 528:2 - 530:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 528:2 - 530:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_27_backwards : string -> M (mword ty27)\\ \ \\val hex_bits_27_forwards_matches : mword ty27 -> bool\\ \ \\val hex_bits_27_backwards_matches : string -> bool\\ definition hex_bits_27_forwards_matches :: \(27)Word.word \ bool \ where \ hex_bits_27_forwards_matches bv = ( True )\ for bv :: "(27)Word.word " definition hex_bits_27_backwards_matches :: \ string \ bool \ where \ hex_bits_27_backwards_matches s = ( if ((case ((hex_bits_27_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 27 Word.word * ii))option)) of Some ((g__336, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_27_backwards :: \ string \((register_value),((27)Word.word),(exception))monad \ where \ hex_bits_27_backwards s = ( (case ((hex_bits_27_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 27 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 547:2 - 549:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 547:2 - 549:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_28_backwards : string -> M (mword ty28)\\ \ \\val hex_bits_28_forwards_matches : mword ty28 -> bool\\ \ \\val hex_bits_28_backwards_matches : string -> bool\\ definition hex_bits_28_forwards_matches :: \(28)Word.word \ bool \ where \ hex_bits_28_forwards_matches bv = ( True )\ for bv :: "(28)Word.word " definition hex_bits_28_backwards_matches :: \ string \ bool \ where \ hex_bits_28_backwards_matches s = ( if ((case ((hex_bits_28_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 28 Word.word * ii))option)) of Some ((g__335, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_28_backwards :: \ string \((register_value),((28)Word.word),(exception))monad \ where \ hex_bits_28_backwards s = ( (case ((hex_bits_28_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 28 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 566:2 - 568:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 566:2 - 568:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_29_backwards : string -> M (mword ty29)\\ \ \\val hex_bits_29_forwards_matches : mword ty29 -> bool\\ \ \\val hex_bits_29_backwards_matches : string -> bool\\ definition hex_bits_29_forwards_matches :: \(29)Word.word \ bool \ where \ hex_bits_29_forwards_matches bv = ( True )\ for bv :: "(29)Word.word " definition hex_bits_29_backwards_matches :: \ string \ bool \ where \ hex_bits_29_backwards_matches s = ( if ((case ((hex_bits_29_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 29 Word.word * ii))option)) of Some ((g__334, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_29_backwards :: \ string \((register_value),((29)Word.word),(exception))monad \ where \ hex_bits_29_backwards s = ( (case ((hex_bits_29_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 29 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 585:2 - 587:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 585:2 - 587:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_30_backwards : string -> M (mword ty30)\\ \ \\val hex_bits_30_forwards_matches : mword ty30 -> bool\\ \ \\val hex_bits_30_backwards_matches : string -> bool\\ definition hex_bits_30_forwards_matches :: \(30)Word.word \ bool \ where \ hex_bits_30_forwards_matches bv = ( True )\ for bv :: "(30)Word.word " definition hex_bits_30_backwards_matches :: \ string \ bool \ where \ hex_bits_30_backwards_matches s = ( if ((case ((hex_bits_30_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 30 Word.word * ii))option)) of Some ((g__333, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_30_backwards :: \ string \((register_value),((30)Word.word),(exception))monad \ where \ hex_bits_30_backwards s = ( (case ((hex_bits_30_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 30 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 604:2 - 606:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 604:2 - 606:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_31_backwards : string -> M (mword ty31)\\ \ \\val hex_bits_31_forwards_matches : mword ty31 -> bool\\ \ \\val hex_bits_31_backwards_matches : string -> bool\\ definition hex_bits_31_forwards_matches :: \(31)Word.word \ bool \ where \ hex_bits_31_forwards_matches bv = ( True )\ for bv :: "(31)Word.word " definition hex_bits_31_backwards_matches :: \ string \ bool \ where \ hex_bits_31_backwards_matches s = ( if ((case ((hex_bits_31_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 31 Word.word * ii))option)) of Some ((g__332, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_31_backwards :: \ string \((register_value),((31)Word.word),(exception))monad \ where \ hex_bits_31_backwards s = ( (case ((hex_bits_31_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 31 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 623:2 - 625:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 623:2 - 625:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_32_backwards : string -> M (mword ty32)\\ \ \\val hex_bits_32_forwards_matches : mword ty32 -> bool\\ \ \\val hex_bits_32_backwards_matches : string -> bool\\ definition hex_bits_32_forwards_matches :: \(32)Word.word \ bool \ where \ hex_bits_32_forwards_matches bv = ( True )\ for bv :: "(32)Word.word " definition hex_bits_32_backwards_matches :: \ string \ bool \ where \ hex_bits_32_backwards_matches s = ( if ((case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 32 Word.word * ii))option)) of Some ((g__331, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_32_backwards :: \ string \((register_value),((32)Word.word),(exception))monad \ where \ hex_bits_32_backwards s = ( (case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 32 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 642:2 - 644:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 642:2 - 644:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_33_backwards : string -> M (mword ty33)\\ \ \\val hex_bits_33_forwards_matches : mword ty33 -> bool\\ \ \\val hex_bits_33_backwards_matches : string -> bool\\ definition hex_bits_33_forwards_matches :: \(33)Word.word \ bool \ where \ hex_bits_33_forwards_matches bv = ( True )\ for bv :: "(33)Word.word " definition hex_bits_33_backwards_matches :: \ string \ bool \ where \ hex_bits_33_backwards_matches s = ( if ((case ((hex_bits_33_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 33 Word.word * ii))option)) of Some ((g__330, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_33_backwards :: \ string \((register_value),((33)Word.word),(exception))monad \ where \ hex_bits_33_backwards s = ( (case ((hex_bits_33_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 33 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 661:2 - 663:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 661:2 - 663:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_48_backwards : string -> M (mword ty48)\\ \ \\val hex_bits_48_forwards_matches : mword ty48 -> bool\\ \ \\val hex_bits_48_backwards_matches : string -> bool\\ definition hex_bits_48_forwards_matches :: \(48)Word.word \ bool \ where \ hex_bits_48_forwards_matches bv = ( True )\ for bv :: "(48)Word.word " definition hex_bits_48_backwards_matches :: \ string \ bool \ where \ hex_bits_48_backwards_matches s = ( if ((case ((hex_bits_48_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 48 Word.word * ii))option)) of Some ((g__329, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_48_backwards :: \ string \((register_value),((48)Word.word),(exception))monad \ where \ hex_bits_48_backwards s = ( (case ((hex_bits_48_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 48 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 680:2 - 682:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 680:2 - 682:3'') \ exit0 () ))\ for s :: " string " \ \\val hex_bits_64_backwards : string -> M (mword ty64)\\ \ \\val hex_bits_64_forwards_matches : mword ty64 -> bool\\ \ \\val hex_bits_64_backwards_matches : string -> bool\\ definition hex_bits_64_forwards_matches :: \(64)Word.word \ bool \ where \ hex_bits_64_forwards_matches bv = ( True )\ for bv :: "(64)Word.word " definition hex_bits_64_backwards_matches :: \ string \ bool \ where \ hex_bits_64_backwards_matches s = ( if ((case ((hex_bits_64_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 64 Word.word * ii))option)) of Some ((g__328, n)) => if (((n = ((string_length s))))) then True else False | _ => False )) then True else False )\ for s :: " string " definition hex_bits_64_backwards :: \ string \((register_value),((64)Word.word),(exception))monad \ where \ hex_bits_64_backwards s = ( (case ((hex_bits_64_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s :: (( 64 Word.word * ii))option)) of Some ((bv, n)) => if (((n = ((string_length s))))) then return bv else assert_exp False (''Pattern match failure at model/prelude_mapping.sail 699:2 - 701:3'') \ exit0 () | _ => assert_exp False (''Pattern match failure at model/prelude_mapping.sail 699:2 - 701:3'') \ exit0 () ))\ for s :: " string " definition default_meta :: \ unit \ where \ default_meta = ( () )\ \ \\val __WriteRAM_Meta : mword ty32 -> integer -> unit -> M unit\\ definition WriteRAM_Meta :: \(32)Word.word \ int \ unit \((register_value),(unit),(exception))monad \ where \ WriteRAM_Meta addr width meta = ( return () )\ for addr :: "(32)Word.word " and width :: " int " and meta :: " unit " \ \\val __ReadRAM_Meta : mword ty32 -> integer -> M unit\\ definition ReadRAM_Meta :: \(32)Word.word \ int \((register_value),(unit),(exception))monad \ where \ ReadRAM_Meta addr width = ( return () )\ for addr :: "(32)Word.word " and width :: " int " \ \\val write_ram_ea : write_kind -> mword ty32 -> integer -> M unit\\ definition write_ram_ea :: \ write_kind \(32)Word.word \ int \((register_value),(unit),(exception))monad \ where \ write_ram_ea wk addr width = ( write_mem_ea instance_Sail2_values_Bitvector_Machine_word_mword_dict wk (( 32 :: int)::ii) addr width )\ for wk :: " write_kind " and addr :: "(32)Word.word " and width :: " int " \ \\val __TraceMemoryWrite : forall 'm 'int8_times_n. Size 'm, Size 'int8_times_n => integer -> mword 'm -> mword 'int8_times_n -> unit\\ \ \\val __TraceMemoryRead : forall 'm 'int8_times_n. Size 'm, Size 'int8_times_n => integer -> mword 'm -> mword 'int8_times_n -> unit\\ definition init_ext_ptw :: \ unit \ where \ init_ext_ptw = ( () )\ \ \\val ext_translate_exception : unit -> unit\\ definition ext_translate_exception :: \ unit \ unit \ where \ ext_translate_exception e = ( e )\ for e :: " unit " \ \\val ext_exc_type_to_bits : unit -> mword ty8\\ definition ext_exc_type_to_bits :: \ unit \(8)Word.word \ where \ ext_exc_type_to_bits e = ( ( 0x18 :: 8 Word.word))\ for e :: " unit " \ \\val num_of_ext_exc_type : unit -> integer\\ definition num_of_ext_exc_type :: \ unit \ int \ where \ num_of_ext_exc_type e = ( (( 24 :: int)::ii))\ for e :: " unit " \ \\val ext_exc_type_to_str : unit -> string\\ definition ext_exc_type_to_str :: \ unit \ string \ where \ ext_exc_type_to_str e = ( (''extension-exception''))\ for e :: " unit " definition xlen_val :: \ int \ where \ xlen_val = ( (( 32 :: int)::ii))\ definition xlen_max_unsigned :: \ int \ where \ xlen_max_unsigned = ( (( 4294967295 :: int)::ii))\ definition xlen_max_signed :: \ int \ where \ xlen_max_signed = ( (( 2147483647 :: int)::ii))\ definition xlen_min_signed :: \ int \ where \ xlen_min_signed = ( ((( 0 :: int)-( 2147483648 :: int))::ii))\ \ \\val regidx_to_regno : mword ty5 -> integer\\ definition regidx_to_regno :: \(5)Word.word \ int \ where \ regidx_to_regno b = ( (let r = (Word.uint b) in r))\ for b :: "(5)Word.word " \ \\val creg2reg_idx : mword ty3 -> mword ty5\\ definition creg2reg_idx :: \(3)Word.word \(5)Word.word \ where \ creg2reg_idx creg = ( (concat_vec ( 0b01 :: 2 Word.word) creg :: 5 Word.word))\ for creg :: "(3)Word.word " definition zreg :: \(5)Word.word \ where \ zreg = ( ( 0b00000 :: 5 Word.word))\ definition ra :: \(5)Word.word \ where \ ra = ( ( 0b00001 :: 5 Word.word))\ definition sp :: \(5)Word.word \ where \ sp = ( ( 0b00010 :: 5 Word.word))\ \ \\val Architecture_of_num : integer -> Architecture\\ definition Architecture_of_num :: \ int \ Architecture \ where \ Architecture_of_num arg1 = ( (let l__269 = arg1 in if (((l__269 = (( 0 :: int)::ii)))) then RV32 else if (((l__269 = (( 1 :: int)::ii)))) then RV64 else RV128))\ for arg1 :: " int " \ \\val num_of_Architecture : Architecture -> integer\\ fun num_of_Architecture :: \ Architecture \ int \ where \ num_of_Architecture RV32 = ( (( 0 :: int)::ii))\ |\ num_of_Architecture RV64 = ( (( 1 :: int)::ii))\ |\ num_of_Architecture RV128 = ( (( 2 :: int)::ii))\ \ \\val architecture : mword ty2 -> maybe Architecture\\ definition architecture :: \(2)Word.word \(Architecture)option \ where \ architecture a = ( (let b__0 = a in if (((b__0 = ( 0b01 :: 2 Word.word)))) then Some RV32 else if (((b__0 = ( 0b10 :: 2 Word.word)))) then Some RV64 else if (((b__0 = ( 0b11 :: 2 Word.word)))) then Some RV128 else None))\ for a :: "(2)Word.word " \ \\val arch_to_bits : Architecture -> mword ty2\\ fun arch_to_bits :: \ Architecture \(2)Word.word \ where \ arch_to_bits RV32 = ( ( 0b01 :: 2 Word.word))\ |\ arch_to_bits RV64 = ( ( 0b10 :: 2 Word.word))\ |\ arch_to_bits RV128 = ( ( 0b11 :: 2 Word.word))\ \ \\val Privilege_of_num : integer -> Privilege\\ definition Privilege_of_num :: \ int \ Privilege \ where \ Privilege_of_num arg1 = ( (let l__267 = arg1 in if (((l__267 = (( 0 :: int)::ii)))) then User else if (((l__267 = (( 1 :: int)::ii)))) then Supervisor else Machine))\ for arg1 :: " int " \ \\val num_of_Privilege : Privilege -> integer\\ fun num_of_Privilege :: \ Privilege \ int \ where \ num_of_Privilege User = ( (( 0 :: int)::ii))\ |\ num_of_Privilege Supervisor = ( (( 1 :: int)::ii))\ |\ num_of_Privilege Machine = ( (( 2 :: int)::ii))\ \ \\val privLevel_to_bits : Privilege -> mword ty2\\ fun privLevel_to_bits :: \ Privilege \(2)Word.word \ where \ privLevel_to_bits User = ( ( 0b00 :: 2 Word.word))\ |\ privLevel_to_bits Supervisor = ( ( 0b01 :: 2 Word.word))\ |\ privLevel_to_bits Machine = ( ( 0b11 :: 2 Word.word))\ \ \\val privLevel_of_bits : mword ty2 -> M Privilege\\ definition privLevel_of_bits :: \(2)Word.word \((register_value),(Privilege),(exception))monad \ where \ privLevel_of_bits p = ( (let b__0 = p in if (((b__0 = ( 0b00 :: 2 Word.word)))) then return User else if (((b__0 = ( 0b01 :: 2 Word.word)))) then return Supervisor else if (((b__0 = ( 0b11 :: 2 Word.word)))) then return Machine else assert_exp False (''Pattern match failure at model/riscv_types.sail 78:2 - 82:3'') \ exit0 () ))\ for p :: "(2)Word.word " \ \\val privLevel_to_str : Privilege -> string\\ fun privLevel_to_str :: \ Privilege \ string \ where \ privLevel_to_str User = ( (''U''))\ |\ privLevel_to_str Supervisor = ( (''S''))\ |\ privLevel_to_str Machine = ( (''M''))\ \ \\val accessType_to_str : AccessType unit -> string\\ \ \\val csr_name : mword ty12 -> string\\ \ \\val exceptionType_to_str : ExceptionType -> string\\ \ \\val print_insn : ast -> M string\\ \ \\val ptw_error_to_str : PTW_Error -> string\\ \ \\val reg_name_abi : mword ty5 -> M string\\ \ \\val Retired_of_num : integer -> Retired\\ definition Retired_of_num :: \ int \ Retired \ where \ Retired_of_num arg1 = ( (let l__266 = arg1 in if (((l__266 = (( 0 :: int)::ii)))) then RETIRE_SUCCESS else RETIRE_FAIL))\ for arg1 :: " int " \ \\val num_of_Retired : Retired -> integer\\ fun num_of_Retired :: \ Retired \ int \ where \ num_of_Retired RETIRE_SUCCESS = ( (( 0 :: int)::ii))\ |\ num_of_Retired RETIRE_FAIL = ( (( 1 :: int)::ii))\ \ \\val word_width_of_num : integer -> word_width\\ definition word_width_of_num :: \ int \ word_width \ where \ word_width_of_num arg1 = ( (let l__263 = arg1 in if (((l__263 = (( 0 :: int)::ii)))) then BYTE else if (((l__263 = (( 1 :: int)::ii)))) then HALF else if (((l__263 = (( 2 :: int)::ii)))) then WORD else DOUBLE))\ for arg1 :: " int " \ \\val num_of_word_width : word_width -> integer\\ fun num_of_word_width :: \ word_width \ int \ where \ num_of_word_width BYTE = ( (( 0 :: int)::ii))\ |\ num_of_word_width HALF = ( (( 1 :: int)::ii))\ |\ num_of_word_width WORD = ( (( 2 :: int)::ii))\ |\ num_of_word_width DOUBLE = ( (( 3 :: int)::ii))\ \ \\val InterruptType_of_num : integer -> InterruptType\\ definition InterruptType_of_num :: \ int \ InterruptType \ where \ InterruptType_of_num arg1 = ( (let l__255 = arg1 in if (((l__255 = (( 0 :: int)::ii)))) then I_U_Software else if (((l__255 = (( 1 :: int)::ii)))) then I_S_Software else if (((l__255 = (( 2 :: int)::ii)))) then I_M_Software else if (((l__255 = (( 3 :: int)::ii)))) then I_U_Timer else if (((l__255 = (( 4 :: int)::ii)))) then I_S_Timer else if (((l__255 = (( 5 :: int)::ii)))) then I_M_Timer else if (((l__255 = (( 6 :: int)::ii)))) then I_U_External else if (((l__255 = (( 7 :: int)::ii)))) then I_S_External else I_M_External))\ for arg1 :: " int " \ \\val num_of_InterruptType : InterruptType -> integer\\ fun num_of_InterruptType :: \ InterruptType \ int \ where \ num_of_InterruptType I_U_Software = ( (( 0 :: int)::ii))\ |\ num_of_InterruptType I_S_Software = ( (( 1 :: int)::ii))\ |\ num_of_InterruptType I_M_Software = ( (( 2 :: int)::ii))\ |\ num_of_InterruptType I_U_Timer = ( (( 3 :: int)::ii))\ |\ num_of_InterruptType I_S_Timer = ( (( 4 :: int)::ii))\ |\ num_of_InterruptType I_M_Timer = ( (( 5 :: int)::ii))\ |\ num_of_InterruptType I_U_External = ( (( 6 :: int)::ii))\ |\ num_of_InterruptType I_S_External = ( (( 7 :: int)::ii))\ |\ num_of_InterruptType I_M_External = ( (( 8 :: int)::ii))\ \ \\val interruptType_to_bits : InterruptType -> mword ty8\\ fun interruptType_to_bits :: \ InterruptType \(8)Word.word \ where \ interruptType_to_bits I_U_Software = ( ( 0x00 :: 8 Word.word))\ |\ interruptType_to_bits I_S_Software = ( ( 0x01 :: 8 Word.word))\ |\ interruptType_to_bits I_M_Software = ( ( 0x03 :: 8 Word.word))\ |\ interruptType_to_bits I_U_Timer = ( ( 0x04 :: 8 Word.word))\ |\ interruptType_to_bits I_S_Timer = ( ( 0x05 :: 8 Word.word))\ |\ interruptType_to_bits I_M_Timer = ( ( 0x07 :: 8 Word.word))\ |\ interruptType_to_bits I_U_External = ( ( 0x08 :: 8 Word.word))\ |\ interruptType_to_bits I_S_External = ( ( 0x09 :: 8 Word.word))\ |\ interruptType_to_bits I_M_External = ( ( 0x0B :: 8 Word.word))\ \ \\val exceptionType_to_bits : ExceptionType -> mword ty8\\ fun exceptionType_to_bits :: \ ExceptionType \(8)Word.word \ where \ exceptionType_to_bits (E_Fetch_Addr_Align (_)) = ( ( 0x00 :: 8 Word.word))\ |\ exceptionType_to_bits (E_Fetch_Access_Fault (_)) = ( ( 0x01 :: 8 Word.word))\ |\ exceptionType_to_bits (E_Illegal_Instr (_)) = ( ( 0x02 :: 8 Word.word))\ |\ exceptionType_to_bits (E_Breakpoint (_)) = ( ( 0x03 :: 8 Word.word))\ |\ exceptionType_to_bits (E_Load_Addr_Align (_)) = ( ( 0x04 :: 8 Word.word))\ |\ exceptionType_to_bits (E_Load_Access_Fault (_)) = ( ( 0x05 :: 8 Word.word))\ |\ exceptionType_to_bits (E_SAMO_Addr_Align (_)) = ( ( 0x06 :: 8 Word.word))\ |\ exceptionType_to_bits (E_SAMO_Access_Fault (_)) = ( ( 0x07 :: 8 Word.word))\ |\ exceptionType_to_bits (E_U_EnvCall (_)) = ( ( 0x08 :: 8 Word.word))\ |\ exceptionType_to_bits (E_S_EnvCall (_)) = ( ( 0x09 :: 8 Word.word))\ |\ exceptionType_to_bits (E_Reserved_10 (_)) = ( ( 0x0A :: 8 Word.word))\ |\ exceptionType_to_bits (E_M_EnvCall (_)) = ( ( 0x0B :: 8 Word.word))\ |\ exceptionType_to_bits (E_Fetch_Page_Fault (_)) = ( ( 0x0C :: 8 Word.word))\ |\ exceptionType_to_bits (E_Load_Page_Fault (_)) = ( ( 0x0D :: 8 Word.word))\ |\ exceptionType_to_bits (E_Reserved_14 (_)) = ( ( 0x0E :: 8 Word.word))\ |\ exceptionType_to_bits (E_SAMO_Page_Fault (_)) = ( ( 0x0F :: 8 Word.word))\ |\ exceptionType_to_bits (E_Extension (e)) = ( (ext_exc_type_to_bits e :: 8 Word.word))\ for e :: " unit " \ \\val num_of_ExceptionType : ExceptionType -> integer\\ fun num_of_ExceptionType :: \ ExceptionType \ int \ where \ num_of_ExceptionType (E_Fetch_Addr_Align (_)) = ( (( 0 :: int)::ii))\ |\ num_of_ExceptionType (E_Fetch_Access_Fault (_)) = ( (( 1 :: int)::ii))\ |\ num_of_ExceptionType (E_Illegal_Instr (_)) = ( (( 2 :: int)::ii))\ |\ num_of_ExceptionType (E_Breakpoint (_)) = ( (( 3 :: int)::ii))\ |\ num_of_ExceptionType (E_Load_Addr_Align (_)) = ( (( 4 :: int)::ii))\ |\ num_of_ExceptionType (E_Load_Access_Fault (_)) = ( (( 5 :: int)::ii))\ |\ num_of_ExceptionType (E_SAMO_Addr_Align (_)) = ( (( 6 :: int)::ii))\ |\ num_of_ExceptionType (E_SAMO_Access_Fault (_)) = ( (( 7 :: int)::ii))\ |\ num_of_ExceptionType (E_U_EnvCall (_)) = ( (( 8 :: int)::ii))\ |\ num_of_ExceptionType (E_S_EnvCall (_)) = ( (( 9 :: int)::ii))\ |\ num_of_ExceptionType (E_Reserved_10 (_)) = ( (( 10 :: int)::ii))\ |\ num_of_ExceptionType (E_M_EnvCall (_)) = ( (( 11 :: int)::ii))\ |\ num_of_ExceptionType (E_Fetch_Page_Fault (_)) = ( (( 12 :: int)::ii))\ |\ num_of_ExceptionType (E_Load_Page_Fault (_)) = ( (( 13 :: int)::ii))\ |\ num_of_ExceptionType (E_Reserved_14 (_)) = ( (( 14 :: int)::ii))\ |\ num_of_ExceptionType (E_SAMO_Page_Fault (_)) = ( (( 15 :: int)::ii))\ |\ num_of_ExceptionType (E_Extension (e)) = ( num_of_ext_exc_type e )\ for e :: " unit " fun exceptionType_to_str :: \ ExceptionType \ string \ where \ exceptionType_to_str (E_Fetch_Addr_Align (_)) = ( (''misaligned-fetch''))\ |\ exceptionType_to_str (E_Fetch_Access_Fault (_)) = ( (''fetch-access-fault''))\ |\ exceptionType_to_str (E_Illegal_Instr (_)) = ( (''illegal-instruction''))\ |\ exceptionType_to_str (E_Breakpoint (_)) = ( (''breakpoint''))\ |\ exceptionType_to_str (E_Load_Addr_Align (_)) = ( (''misaligned-load''))\ |\ exceptionType_to_str (E_Load_Access_Fault (_)) = ( (''load-access-fault''))\ |\ exceptionType_to_str (E_SAMO_Addr_Align (_)) = ( (''misaliged-store/amo''))\ |\ exceptionType_to_str (E_SAMO_Access_Fault (_)) = ( (''store/amo-access-fault''))\ |\ exceptionType_to_str (E_U_EnvCall (_)) = ( (''u-call''))\ |\ exceptionType_to_str (E_S_EnvCall (_)) = ( (''s-call''))\ |\ exceptionType_to_str (E_Reserved_10 (_)) = ( (''reserved-0''))\ |\ exceptionType_to_str (E_M_EnvCall (_)) = ( (''m-call''))\ |\ exceptionType_to_str (E_Fetch_Page_Fault (_)) = ( (''fetch-page-fault''))\ |\ exceptionType_to_str (E_Load_Page_Fault (_)) = ( (''load-page-fault''))\ |\ exceptionType_to_str (E_Reserved_14 (_)) = ( (''reserved-1''))\ |\ exceptionType_to_str (E_SAMO_Page_Fault (_)) = ( (''store/amo-page-fault''))\ |\ exceptionType_to_str (E_Extension (e)) = ( ext_exc_type_to_str e )\ for e :: " unit " \ \\val not_implemented : forall 'a. string -> M 'a\\ definition not_implemented :: \ string \((register_value),'a,(exception))monad \ where \ not_implemented message = ( throw (Error_not_implemented message))\ for message :: " string " \ \\val internal_error : forall 'a. string -> M 'a\\ definition internal_error :: \ string \((register_value),'a,(exception))monad \ where \ internal_error s = ( assert_exp False s \ exit0 () )\ for s :: " string " \ \\val TrapVectorMode_of_num : integer -> TrapVectorMode\\ definition TrapVectorMode_of_num :: \ int \ TrapVectorMode \ where \ TrapVectorMode_of_num arg1 = ( (let l__253 = arg1 in if (((l__253 = (( 0 :: int)::ii)))) then TV_Direct else if (((l__253 = (( 1 :: int)::ii)))) then TV_Vector else TV_Reserved))\ for arg1 :: " int " \ \\val num_of_TrapVectorMode : TrapVectorMode -> integer\\ fun num_of_TrapVectorMode :: \ TrapVectorMode \ int \ where \ num_of_TrapVectorMode TV_Direct = ( (( 0 :: int)::ii))\ |\ num_of_TrapVectorMode TV_Vector = ( (( 1 :: int)::ii))\ |\ num_of_TrapVectorMode TV_Reserved = ( (( 2 :: int)::ii))\ \ \\val trapVectorMode_of_bits : mword ty2 -> TrapVectorMode\\ definition trapVectorMode_of_bits :: \(2)Word.word \ TrapVectorMode \ where \ trapVectorMode_of_bits m = ( (let b__0 = m in if (((b__0 = ( 0b00 :: 2 Word.word)))) then TV_Direct else if (((b__0 = ( 0b01 :: 2 Word.word)))) then TV_Vector else TV_Reserved))\ for m :: "(2)Word.word " \ \\val ExtStatus_of_num : integer -> ExtStatus\\ definition ExtStatus_of_num :: \ int \ ExtStatus \ where \ ExtStatus_of_num arg1 = ( (let l__250 = arg1 in if (((l__250 = (( 0 :: int)::ii)))) then Off else if (((l__250 = (( 1 :: int)::ii)))) then Initial else if (((l__250 = (( 2 :: int)::ii)))) then Clean else Dirty))\ for arg1 :: " int " \ \\val num_of_ExtStatus : ExtStatus -> integer\\ fun num_of_ExtStatus :: \ ExtStatus \ int \ where \ num_of_ExtStatus Off = ( (( 0 :: int)::ii))\ |\ num_of_ExtStatus Initial = ( (( 1 :: int)::ii))\ |\ num_of_ExtStatus Clean = ( (( 2 :: int)::ii))\ |\ num_of_ExtStatus Dirty = ( (( 3 :: int)::ii))\ \ \\val extStatus_to_bits : ExtStatus -> mword ty2\\ fun extStatus_to_bits :: \ ExtStatus \(2)Word.word \ where \ extStatus_to_bits Off = ( ( 0b00 :: 2 Word.word))\ |\ extStatus_to_bits Initial = ( ( 0b01 :: 2 Word.word))\ |\ extStatus_to_bits Clean = ( ( 0b10 :: 2 Word.word))\ |\ extStatus_to_bits Dirty = ( ( 0b11 :: 2 Word.word))\ \ \\val extStatus_of_bits : mword ty2 -> M ExtStatus\\ definition extStatus_of_bits :: \(2)Word.word \((register_value),(ExtStatus),(exception))monad \ where \ extStatus_of_bits e = ( (let b__0 = e in if (((b__0 = ( 0b00 :: 2 Word.word)))) then return Off else if (((b__0 = ( 0b01 :: 2 Word.word)))) then return Initial else if (((b__0 = ( 0b10 :: 2 Word.word)))) then return Clean else if (((b__0 = ( 0b11 :: 2 Word.word)))) then return Dirty else assert_exp False (''Pattern match failure at model/riscv_types.sail 281:2 - 286:3'') \ exit0 () ))\ for e :: "(2)Word.word " \ \\val SATPMode_of_num : integer -> SATPMode\\ definition SATPMode_of_num :: \ int \ SATPMode \ where \ SATPMode_of_num arg1 = ( (let l__247 = arg1 in if (((l__247 = (( 0 :: int)::ii)))) then Sbare else if (((l__247 = (( 1 :: int)::ii)))) then Sv32 else if (((l__247 = (( 2 :: int)::ii)))) then Sv39 else Sv48))\ for arg1 :: " int " \ \\val num_of_SATPMode : SATPMode -> integer\\ fun num_of_SATPMode :: \ SATPMode \ int \ where \ num_of_SATPMode Sbare = ( (( 0 :: int)::ii))\ |\ num_of_SATPMode Sv32 = ( (( 1 :: int)::ii))\ |\ num_of_SATPMode Sv39 = ( (( 2 :: int)::ii))\ |\ num_of_SATPMode Sv48 = ( (( 3 :: int)::ii))\ \ \\val satp64Mode_of_bits : Architecture -> mword ty4 -> maybe SATPMode\\ definition satp64Mode_of_bits :: \ Architecture \(4)Word.word \(SATPMode)option \ where \ satp64Mode_of_bits (g__327 :: Architecture) (b__0 :: satp_mode) = ( if (((b__0 = ( 0x0 :: 4 Word.word)))) then Some Sbare else (case (g__327, b__0) of (RV32, b__0) => if (((b__0 = ( 0x1 :: 4 Word.word)))) then Some Sv32 else (case (RV32, b__0) of (_, _) => None ) | (RV64, b__0) => if (((b__0 = ( 0x8 :: 4 Word.word)))) then Some Sv39 else if (((b__0 = ( 0x9 :: 4 Word.word)))) then Some Sv48 else (case (RV64, b__0) of (_, _) => None ) | (_, _) => None ))\ for g__327 :: " Architecture " and b__0 :: "(4)Word.word " \ \\val uop_of_num : integer -> uop\\ definition uop_of_num :: \ int \ uop \ where \ uop_of_num arg1 = ( (let l__246 = arg1 in if (((l__246 = (( 0 :: int)::ii)))) then RISCV_LUI else RISCV_AUIPC))\ for arg1 :: " int " \ \\val num_of_uop : uop -> integer\\ fun num_of_uop :: \ uop \ int \ where \ num_of_uop RISCV_LUI = ( (( 0 :: int)::ii))\ |\ num_of_uop RISCV_AUIPC = ( (( 1 :: int)::ii))\ \ \\val bop_of_num : integer -> bop\\ definition bop_of_num :: \ int \ bop \ where \ bop_of_num arg1 = ( (let l__241 = arg1 in if (((l__241 = (( 0 :: int)::ii)))) then RISCV_BEQ else if (((l__241 = (( 1 :: int)::ii)))) then RISCV_BNE else if (((l__241 = (( 2 :: int)::ii)))) then RISCV_BLT else if (((l__241 = (( 3 :: int)::ii)))) then RISCV_BGE else if (((l__241 = (( 4 :: int)::ii)))) then RISCV_BLTU else RISCV_BGEU))\ for arg1 :: " int " \ \\val num_of_bop : bop -> integer\\ fun num_of_bop :: \ bop \ int \ where \ num_of_bop RISCV_BEQ = ( (( 0 :: int)::ii))\ |\ num_of_bop RISCV_BNE = ( (( 1 :: int)::ii))\ |\ num_of_bop RISCV_BLT = ( (( 2 :: int)::ii))\ |\ num_of_bop RISCV_BGE = ( (( 3 :: int)::ii))\ |\ num_of_bop RISCV_BLTU = ( (( 4 :: int)::ii))\ |\ num_of_bop RISCV_BGEU = ( (( 5 :: int)::ii))\ \ \\val iop_of_num : integer -> iop\\ definition iop_of_num :: \ int \ iop \ where \ iop_of_num arg1 = ( (let l__236 = arg1 in if (((l__236 = (( 0 :: int)::ii)))) then RISCV_ADDI else if (((l__236 = (( 1 :: int)::ii)))) then RISCV_SLTI else if (((l__236 = (( 2 :: int)::ii)))) then RISCV_SLTIU else if (((l__236 = (( 3 :: int)::ii)))) then RISCV_XORI else if (((l__236 = (( 4 :: int)::ii)))) then RISCV_ORI else RISCV_ANDI))\ for arg1 :: " int " \ \\val num_of_iop : iop -> integer\\ fun num_of_iop :: \ iop \ int \ where \ num_of_iop RISCV_ADDI = ( (( 0 :: int)::ii))\ |\ num_of_iop RISCV_SLTI = ( (( 1 :: int)::ii))\ |\ num_of_iop RISCV_SLTIU = ( (( 2 :: int)::ii))\ |\ num_of_iop RISCV_XORI = ( (( 3 :: int)::ii))\ |\ num_of_iop RISCV_ORI = ( (( 4 :: int)::ii))\ |\ num_of_iop RISCV_ANDI = ( (( 5 :: int)::ii))\ \ \\val sop_of_num : integer -> sop\\ definition sop_of_num :: \ int \ sop \ where \ sop_of_num arg1 = ( (let l__234 = arg1 in if (((l__234 = (( 0 :: int)::ii)))) then RISCV_SLLI else if (((l__234 = (( 1 :: int)::ii)))) then RISCV_SRLI else RISCV_SRAI))\ for arg1 :: " int " \ \\val num_of_sop : sop -> integer\\ fun num_of_sop :: \ sop \ int \ where \ num_of_sop RISCV_SLLI = ( (( 0 :: int)::ii))\ |\ num_of_sop RISCV_SRLI = ( (( 1 :: int)::ii))\ |\ num_of_sop RISCV_SRAI = ( (( 2 :: int)::ii))\ \ \\val rop_of_num : integer -> rop\\ definition rop_of_num :: \ int \ rop \ where \ rop_of_num arg1 = ( (let l__225 = arg1 in if (((l__225 = (( 0 :: int)::ii)))) then RISCV_ADD else if (((l__225 = (( 1 :: int)::ii)))) then RISCV_SUB else if (((l__225 = (( 2 :: int)::ii)))) then RISCV_SLL else if (((l__225 = (( 3 :: int)::ii)))) then RISCV_SLT else if (((l__225 = (( 4 :: int)::ii)))) then RISCV_SLTU else if (((l__225 = (( 5 :: int)::ii)))) then RISCV_XOR else if (((l__225 = (( 6 :: int)::ii)))) then RISCV_SRL else if (((l__225 = (( 7 :: int)::ii)))) then RISCV_SRA else if (((l__225 = (( 8 :: int)::ii)))) then RISCV_OR else RISCV_AND))\ for arg1 :: " int " \ \\val num_of_rop : rop -> integer\\ fun num_of_rop :: \ rop \ int \ where \ num_of_rop RISCV_ADD = ( (( 0 :: int)::ii))\ |\ num_of_rop RISCV_SUB = ( (( 1 :: int)::ii))\ |\ num_of_rop RISCV_SLL = ( (( 2 :: int)::ii))\ |\ num_of_rop RISCV_SLT = ( (( 3 :: int)::ii))\ |\ num_of_rop RISCV_SLTU = ( (( 4 :: int)::ii))\ |\ num_of_rop RISCV_XOR = ( (( 5 :: int)::ii))\ |\ num_of_rop RISCV_SRL = ( (( 6 :: int)::ii))\ |\ num_of_rop RISCV_SRA = ( (( 7 :: int)::ii))\ |\ num_of_rop RISCV_OR = ( (( 8 :: int)::ii))\ |\ num_of_rop RISCV_AND = ( (( 9 :: int)::ii))\ \ \\val ropw_of_num : integer -> ropw\\ definition ropw_of_num :: \ int \ ropw \ where \ ropw_of_num arg1 = ( (let l__221 = arg1 in if (((l__221 = (( 0 :: int)::ii)))) then RISCV_ADDW else if (((l__221 = (( 1 :: int)::ii)))) then RISCV_SUBW else if (((l__221 = (( 2 :: int)::ii)))) then RISCV_SLLW else if (((l__221 = (( 3 :: int)::ii)))) then RISCV_SRLW else RISCV_SRAW))\ for arg1 :: " int " \ \\val num_of_ropw : ropw -> integer\\ fun num_of_ropw :: \ ropw \ int \ where \ num_of_ropw RISCV_ADDW = ( (( 0 :: int)::ii))\ |\ num_of_ropw RISCV_SUBW = ( (( 1 :: int)::ii))\ |\ num_of_ropw RISCV_SLLW = ( (( 2 :: int)::ii))\ |\ num_of_ropw RISCV_SRLW = ( (( 3 :: int)::ii))\ |\ num_of_ropw RISCV_SRAW = ( (( 4 :: int)::ii))\ \ \\val sopw_of_num : integer -> sopw\\ definition sopw_of_num :: \ int \ sopw \ where \ sopw_of_num arg1 = ( (let l__219 = arg1 in if (((l__219 = (( 0 :: int)::ii)))) then RISCV_SLLIW else if (((l__219 = (( 1 :: int)::ii)))) then RISCV_SRLIW else RISCV_SRAIW))\ for arg1 :: " int " \ \\val num_of_sopw : sopw -> integer\\ fun num_of_sopw :: \ sopw \ int \ where \ num_of_sopw RISCV_SLLIW = ( (( 0 :: int)::ii))\ |\ num_of_sopw RISCV_SRLIW = ( (( 1 :: int)::ii))\ |\ num_of_sopw RISCV_SRAIW = ( (( 2 :: int)::ii))\ \ \\val amoop_of_num : integer -> amoop\\ definition amoop_of_num :: \ int \ amoop \ where \ amoop_of_num arg1 = ( (let l__211 = arg1 in if (((l__211 = (( 0 :: int)::ii)))) then AMOSWAP else if (((l__211 = (( 1 :: int)::ii)))) then AMOADD else if (((l__211 = (( 2 :: int)::ii)))) then AMOXOR else if (((l__211 = (( 3 :: int)::ii)))) then AMOAND else if (((l__211 = (( 4 :: int)::ii)))) then AMOOR else if (((l__211 = (( 5 :: int)::ii)))) then AMOMIN else if (((l__211 = (( 6 :: int)::ii)))) then AMOMAX else if (((l__211 = (( 7 :: int)::ii)))) then AMOMINU else AMOMAXU))\ for arg1 :: " int " \ \\val num_of_amoop : amoop -> integer\\ fun num_of_amoop :: \ amoop \ int \ where \ num_of_amoop AMOSWAP = ( (( 0 :: int)::ii))\ |\ num_of_amoop AMOADD = ( (( 1 :: int)::ii))\ |\ num_of_amoop AMOXOR = ( (( 2 :: int)::ii))\ |\ num_of_amoop AMOAND = ( (( 3 :: int)::ii))\ |\ num_of_amoop AMOOR = ( (( 4 :: int)::ii))\ |\ num_of_amoop AMOMIN = ( (( 5 :: int)::ii))\ |\ num_of_amoop AMOMAX = ( (( 6 :: int)::ii))\ |\ num_of_amoop AMOMINU = ( (( 7 :: int)::ii))\ |\ num_of_amoop AMOMAXU = ( (( 8 :: int)::ii))\ \ \\val csrop_of_num : integer -> csrop\\ definition csrop_of_num :: \ int \ csrop \ where \ csrop_of_num arg1 = ( (let l__209 = arg1 in if (((l__209 = (( 0 :: int)::ii)))) then CSRRW else if (((l__209 = (( 1 :: int)::ii)))) then CSRRS else CSRRC))\ for arg1 :: " int " \ \\val num_of_csrop : csrop -> integer\\ fun num_of_csrop :: \ csrop \ int \ where \ num_of_csrop CSRRW = ( (( 0 :: int)::ii))\ |\ num_of_csrop CSRRS = ( (( 1 :: int)::ii))\ |\ num_of_csrop CSRRC = ( (( 2 :: int)::ii))\ \ \\val sep_forwards : unit -> string\\ \ \\val sep_backwards : string -> M unit\\ \ \\val sep_forwards_matches : unit -> bool\\ \ \\val sep_backwards_matches : string -> bool\\ \ \\val sep_matches_prefix : string -> maybe ((unit * ii))\\ definition sep_forwards :: \ unit \ string \ where \ sep_forwards _ = ( string_append ((opt_spc_forwards () )) ((string_append ('','') ((string_append ((def_spc_forwards () )) (''''))))))\ \ \\val _s0_ : string -> maybe unit\\ definition s0 :: \ string \(unit)option \ where \ s0 s20 = ( (case ((opt_spc_matches_prefix0 s20)) of Some ((_, s30)) => (let s41 = (string_drop s20 s30) in if ((string_startswith s41 ('',''))) then (case ((string_drop s41 ((string_length ('',''))))) of s50 => (case ((def_spc_matches_prefix s50)) of Some ((_, s61)) => (let p00 = (string_drop s50 s61) in if (((p00 = ('''')))) then Some () else None) | _ => None ) ) else None) | _ => None ))\ for s20 :: " string " definition sep_backwards :: \ string \((register_value),(unit),(exception))monad \ where \ sep_backwards arg1 = ( (let s70 = arg1 in if ((case ((s0 s70)) of Some (_) => True | _ => False )) then (case s0 s70 of (Some (_)) => return () ) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition sep_forwards_matches :: \ unit \ bool \ where \ sep_forwards_matches _ = ( True )\ \ \\val _s8_ : string -> maybe unit\\ definition s8 :: \ string \(unit)option \ where \ s8 s101 = ( (case ((opt_spc_matches_prefix0 s101)) of Some ((_, s110)) => (let s121 = (string_drop s101 s110) in if ((string_startswith s121 ('',''))) then (case ((string_drop s121 ((string_length ('',''))))) of s130 => (case ((def_spc_matches_prefix s130)) of Some ((_, s141)) => (let p00 = (string_drop s130 s141) in if (((p00 = ('''')))) then Some () else None) | _ => None ) ) else None) | _ => None ))\ for s101 :: " string " definition sep_backwards_matches :: \ string \ bool \ where \ sep_backwards_matches arg1 = ( (let s150 = arg1 in if ((case ((s8 s150)) of Some (_) => True | _ => False )) then (case s8 s150 of (Some (_)) => True ) else False))\ for arg1 :: " string " \ \\val _s16_ : string -> maybe string\\ definition s16 :: \ string \(string)option \ where \ s16 s181 = ( (case ((opt_spc_matches_prefix0 s181)) of Some ((_, s190)) => (let s201 = (string_drop s181 s190) in if ((string_startswith s201 ('',''))) then (case ((string_drop s201 ((string_length ('',''))))) of s210 => (case ((def_spc_matches_prefix s210)) of Some ((_, s221)) => (case ((string_drop s210 s221)) of s1 => Some s1 ) | _ => None ) ) else None) | _ => None ))\ for s181 :: " string " definition sep_matches_prefix :: \ string \(unit*int)option \ where \ sep_matches_prefix arg1 = ( (let s230 = arg1 in if ((case ((s16 s230)) of Some (s1) => True | _ => False )) then (case s16 s230 of (Some (s1)) => Some (() , ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val bool_bits_forwards : bool -> mword ty1\\ \ \\val bool_bits_backwards : mword ty1 -> M bool\\ \ \\val bool_bits_forwards_matches : bool -> bool\\ \ \\val bool_bits_backwards_matches : mword ty1 -> bool\\ fun bool_bits_forwards :: \ bool \(1)Word.word \ where \ bool_bits_forwards True = ( ( 0b1 :: 1 Word.word))\ |\ bool_bits_forwards False = ( ( 0b0 :: 1 Word.word))\ definition bool_bits_backwards :: \(1)Word.word \((register_value),(bool),(exception))monad \ where \ bool_bits_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then return True else if (((b__0 = ( 0b0 :: 1 Word.word)))) then return False else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(1)Word.word " fun bool_bits_forwards_matches :: \ bool \ bool \ where \ bool_bits_forwards_matches True = ( True )\ |\ bool_bits_forwards_matches False = ( True )\ definition bool_bits_backwards_matches :: \(1)Word.word \ bool \ where \ bool_bits_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then True else if (((b__0 = ( 0b0 :: 1 Word.word)))) then True else False))\ for arg1 :: "(1)Word.word " \ \\val bool_not_bits_forwards : bool -> mword ty1\\ \ \\val bool_not_bits_backwards : mword ty1 -> M bool\\ \ \\val bool_not_bits_forwards_matches : bool -> bool\\ \ \\val bool_not_bits_backwards_matches : mword ty1 -> bool\\ fun bool_not_bits_forwards :: \ bool \(1)Word.word \ where \ bool_not_bits_forwards True = ( ( 0b0 :: 1 Word.word))\ |\ bool_not_bits_forwards False = ( ( 0b1 :: 1 Word.word))\ definition bool_not_bits_backwards :: \(1)Word.word \((register_value),(bool),(exception))monad \ where \ bool_not_bits_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b0 :: 1 Word.word)))) then return True else if (((b__0 = ( 0b1 :: 1 Word.word)))) then return False else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(1)Word.word " fun bool_not_bits_forwards_matches :: \ bool \ bool \ where \ bool_not_bits_forwards_matches True = ( True )\ |\ bool_not_bits_forwards_matches False = ( True )\ definition bool_not_bits_backwards_matches :: \(1)Word.word \ bool \ where \ bool_not_bits_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b0 :: 1 Word.word)))) then True else if (((b__0 = ( 0b1 :: 1 Word.word)))) then True else False))\ for arg1 :: "(1)Word.word " \ \\val size_bits_forwards : word_width -> mword ty2\\ \ \\val size_bits_backwards : mword ty2 -> M word_width\\ \ \\val size_bits_forwards_matches : word_width -> bool\\ \ \\val size_bits_backwards_matches : mword ty2 -> bool\\ fun size_bits_forwards :: \ word_width \(2)Word.word \ where \ size_bits_forwards BYTE = ( ( 0b00 :: 2 Word.word))\ |\ size_bits_forwards HALF = ( ( 0b01 :: 2 Word.word))\ |\ size_bits_forwards WORD = ( ( 0b10 :: 2 Word.word))\ |\ size_bits_forwards DOUBLE = ( ( 0b11 :: 2 Word.word))\ definition size_bits_backwards :: \(2)Word.word \((register_value),(word_width),(exception))monad \ where \ size_bits_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00 :: 2 Word.word)))) then return BYTE else if (((b__0 = ( 0b01 :: 2 Word.word)))) then return HALF else if (((b__0 = ( 0b10 :: 2 Word.word)))) then return WORD else if (((b__0 = ( 0b11 :: 2 Word.word)))) then return DOUBLE else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(2)Word.word " fun size_bits_forwards_matches :: \ word_width \ bool \ where \ size_bits_forwards_matches BYTE = ( True )\ |\ size_bits_forwards_matches HALF = ( True )\ |\ size_bits_forwards_matches WORD = ( True )\ |\ size_bits_forwards_matches DOUBLE = ( True )\ definition size_bits_backwards_matches :: \(2)Word.word \ bool \ where \ size_bits_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00 :: 2 Word.word)))) then True else if (((b__0 = ( 0b01 :: 2 Word.word)))) then True else if (((b__0 = ( 0b10 :: 2 Word.word)))) then True else if (((b__0 = ( 0b11 :: 2 Word.word)))) then True else False))\ for arg1 :: "(2)Word.word " \ \\val size_mnemonic_forwards : word_width -> string\\ \ \\val size_mnemonic_backwards : string -> M word_width\\ \ \\val size_mnemonic_forwards_matches : word_width -> bool\\ \ \\val size_mnemonic_backwards_matches : string -> bool\\ \ \\val size_mnemonic_matches_prefix : string -> maybe ((word_width * ii))\\ fun size_mnemonic_forwards :: \ word_width \ string \ where \ size_mnemonic_forwards BYTE = ( (''b''))\ |\ size_mnemonic_forwards HALF = ( (''h''))\ |\ size_mnemonic_forwards WORD = ( (''w''))\ |\ size_mnemonic_forwards DOUBLE = ( (''d''))\ definition size_mnemonic_backwards :: \ string \((register_value),(word_width),(exception))monad \ where \ size_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''b'')))) then return BYTE else if (((p00 = (''h'')))) then return HALF else if (((p00 = (''w'')))) then return WORD else if (((p00 = (''d'')))) then return DOUBLE else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun size_mnemonic_forwards_matches :: \ word_width \ bool \ where \ size_mnemonic_forwards_matches BYTE = ( True )\ |\ size_mnemonic_forwards_matches HALF = ( True )\ |\ size_mnemonic_forwards_matches WORD = ( True )\ |\ size_mnemonic_forwards_matches DOUBLE = ( True )\ definition size_mnemonic_backwards_matches :: \ string \ bool \ where \ size_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''b'')))) then True else if (((p00 = (''h'')))) then True else if (((p00 = (''w'')))) then True else if (((p00 = (''d'')))) then True else False))\ for arg1 :: " string " \ \\val _s36_ : string -> maybe string\\ definition s36 :: \ string \(string)option \ where \ s36 s370 = ( (let s381 = s370 in if ((string_startswith s381 (''d''))) then (case ((string_drop s381 ((string_length (''d''))))) of s1 => Some s1 ) else None))\ for s370 :: " string " \ \\val _s32_ : string -> maybe string\\ definition s32 :: \ string \(string)option \ where \ s32 s330 = ( (let s341 = s330 in if ((string_startswith s341 (''w''))) then (case ((string_drop s341 ((string_length (''w''))))) of s1 => Some s1 ) else None))\ for s330 :: " string " \ \\val _s28_ : string -> maybe string\\ definition s28 :: \ string \(string)option \ where \ s28 s290 = ( (let s301 = s290 in if ((string_startswith s301 (''h''))) then (case ((string_drop s301 ((string_length (''h''))))) of s1 => Some s1 ) else None))\ for s290 :: " string " \ \\val _s24_ : string -> maybe string\\ definition s24 :: \ string \(string)option \ where \ s24 s250 = ( (let s261 = s250 in if ((string_startswith s261 (''b''))) then (case ((string_drop s261 ((string_length (''b''))))) of s1 => Some s1 ) else None))\ for s250 :: " string " definition size_mnemonic_matches_prefix :: \ string \(word_width*int)option \ where \ size_mnemonic_matches_prefix arg1 = ( (let s270 = arg1 in if ((case ((s24 s270)) of Some (s1) => True | _ => False )) then (case s24 s270 of (Some (s1)) => Some (BYTE, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s28 s270)) of Some (s1) => True | _ => False )) then (case s28 s270 of (Some (s1)) => Some (HALF, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s32 s270)) of Some (s1) => True | _ => False )) then (case s32 s270 of (Some (s1)) => Some (WORD, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s36 s270)) of Some (s1) => True | _ => False )) then (case s36 s270 of (Some (s1)) => Some (DOUBLE, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val word_width_bytes : word_width -> integer\\ fun word_width_bytes :: \ word_width \ int \ where \ word_width_bytes BYTE = ( (( 1 :: int)::ii))\ |\ word_width_bytes HALF = ( (( 2 :: int)::ii))\ |\ word_width_bytes WORD = ( (( 4 :: int)::ii))\ |\ word_width_bytes DOUBLE = ( (( 8 :: int)::ii))\ definition Data :: \ unit \ where \ Data = ( () )\ definition default_write_acc :: \ unit \ where \ default_write_acc = ( () )\ fun accessType_to_str :: \(unit)AccessType \ string \ where \ accessType_to_str (Read (_)) = ( (''R''))\ |\ accessType_to_str (Write (_)) = ( (''W''))\ |\ accessType_to_str (ReadWrite ((_, _))) = ( (''RW''))\ |\ accessType_to_str (Execute (_)) = ( (''X''))\ definition zero_reg :: \(32)Word.word \ where \ zero_reg = ( ( 0x00000000 :: 32 Word.word))\ \ \\val RegStr : mword ty32 -> string\\ definition RegStr :: \(32)Word.word \ string \ where \ RegStr r = ( string_of_bits r )\ for r :: "(32)Word.word " \ \\val regval_from_reg : mword ty32 -> mword ty32\\ definition regval_from_reg :: \(32)Word.word \(32)Word.word \ where \ regval_from_reg r = ( r )\ for r :: "(32)Word.word " \ \\val regval_into_reg : mword ty32 -> mword ty32\\ definition regval_into_reg :: \(32)Word.word \(32)Word.word \ where \ regval_into_reg v = ( v )\ for v :: "(32)Word.word " definition zero_freg :: \(32)Word.word \ where \ zero_freg = ( ( 0x00000000 :: 32 Word.word))\ \ \\val FRegStr : mword ty32 -> string\\ definition FRegStr :: \(32)Word.word \ string \ where \ FRegStr r = ( string_of_bits r )\ for r :: "(32)Word.word " \ \\val fregval_from_freg : mword ty32 -> mword ty32\\ definition fregval_from_freg :: \(32)Word.word \(32)Word.word \ where \ fregval_from_freg r = ( r )\ for r :: "(32)Word.word " \ \\val fregval_into_freg : mword ty32 -> mword ty32\\ definition fregval_into_freg :: \(32)Word.word \(32)Word.word \ where \ fregval_into_freg v = ( v )\ for v :: "(32)Word.word " \ \\val rounding_mode_of_num : integer -> rounding_mode\\ definition rounding_mode_of_num :: \ int \ rounding_mode \ where \ rounding_mode_of_num arg1 = ( (let l__204 = arg1 in if (((l__204 = (( 0 :: int)::ii)))) then RM_RNE else if (((l__204 = (( 1 :: int)::ii)))) then RM_RTZ else if (((l__204 = (( 2 :: int)::ii)))) then RM_RDN else if (((l__204 = (( 3 :: int)::ii)))) then RM_RUP else if (((l__204 = (( 4 :: int)::ii)))) then RM_RMM else RM_DYN))\ for arg1 :: " int " \ \\val num_of_rounding_mode : rounding_mode -> integer\\ fun num_of_rounding_mode :: \ rounding_mode \ int \ where \ num_of_rounding_mode RM_RNE = ( (( 0 :: int)::ii))\ |\ num_of_rounding_mode RM_RTZ = ( (( 1 :: int)::ii))\ |\ num_of_rounding_mode RM_RDN = ( (( 2 :: int)::ii))\ |\ num_of_rounding_mode RM_RUP = ( (( 3 :: int)::ii))\ |\ num_of_rounding_mode RM_RMM = ( (( 4 :: int)::ii))\ |\ num_of_rounding_mode RM_DYN = ( (( 5 :: int)::ii))\ \ \\val f_madd_op_S_of_num : integer -> f_madd_op_S\\ definition f_madd_op_S_of_num :: \ int \ f_madd_op_S \ where \ f_madd_op_S_of_num arg1 = ( (let l__201 = arg1 in if (((l__201 = (( 0 :: int)::ii)))) then FMADD_S else if (((l__201 = (( 1 :: int)::ii)))) then FMSUB_S else if (((l__201 = (( 2 :: int)::ii)))) then FNMSUB_S else FNMADD_S))\ for arg1 :: " int " \ \\val num_of_f_madd_op_S : f_madd_op_S -> integer\\ fun num_of_f_madd_op_S :: \ f_madd_op_S \ int \ where \ num_of_f_madd_op_S FMADD_S = ( (( 0 :: int)::ii))\ |\ num_of_f_madd_op_S FMSUB_S = ( (( 1 :: int)::ii))\ |\ num_of_f_madd_op_S FNMSUB_S = ( (( 2 :: int)::ii))\ |\ num_of_f_madd_op_S FNMADD_S = ( (( 3 :: int)::ii))\ \ \\val f_bin_rm_op_S_of_num : integer -> f_bin_rm_op_S\\ definition f_bin_rm_op_S_of_num :: \ int \ f_bin_rm_op_S \ where \ f_bin_rm_op_S_of_num arg1 = ( (let l__198 = arg1 in if (((l__198 = (( 0 :: int)::ii)))) then FADD_S else if (((l__198 = (( 1 :: int)::ii)))) then FSUB_S else if (((l__198 = (( 2 :: int)::ii)))) then FMUL_S else FDIV_S))\ for arg1 :: " int " \ \\val num_of_f_bin_rm_op_S : f_bin_rm_op_S -> integer\\ fun num_of_f_bin_rm_op_S :: \ f_bin_rm_op_S \ int \ where \ num_of_f_bin_rm_op_S FADD_S = ( (( 0 :: int)::ii))\ |\ num_of_f_bin_rm_op_S FSUB_S = ( (( 1 :: int)::ii))\ |\ num_of_f_bin_rm_op_S FMUL_S = ( (( 2 :: int)::ii))\ |\ num_of_f_bin_rm_op_S FDIV_S = ( (( 3 :: int)::ii))\ \ \\val f_un_rm_op_S_of_num : integer -> f_un_rm_op_S\\ definition f_un_rm_op_S_of_num :: \ int \ f_un_rm_op_S \ where \ f_un_rm_op_S_of_num arg1 = ( (let l__190 = arg1 in if (((l__190 = (( 0 :: int)::ii)))) then FSQRT_S else if (((l__190 = (( 1 :: int)::ii)))) then FCVT_W_S else if (((l__190 = (( 2 :: int)::ii)))) then FCVT_WU_S else if (((l__190 = (( 3 :: int)::ii)))) then FCVT_S_W else if (((l__190 = (( 4 :: int)::ii)))) then FCVT_S_WU else if (((l__190 = (( 5 :: int)::ii)))) then FCVT_L_S else if (((l__190 = (( 6 :: int)::ii)))) then FCVT_LU_S else if (((l__190 = (( 7 :: int)::ii)))) then FCVT_S_L else FCVT_S_LU))\ for arg1 :: " int " \ \\val num_of_f_un_rm_op_S : f_un_rm_op_S -> integer\\ fun num_of_f_un_rm_op_S :: \ f_un_rm_op_S \ int \ where \ num_of_f_un_rm_op_S FSQRT_S = ( (( 0 :: int)::ii))\ |\ num_of_f_un_rm_op_S FCVT_W_S = ( (( 1 :: int)::ii))\ |\ num_of_f_un_rm_op_S FCVT_WU_S = ( (( 2 :: int)::ii))\ |\ num_of_f_un_rm_op_S FCVT_S_W = ( (( 3 :: int)::ii))\ |\ num_of_f_un_rm_op_S FCVT_S_WU = ( (( 4 :: int)::ii))\ |\ num_of_f_un_rm_op_S FCVT_L_S = ( (( 5 :: int)::ii))\ |\ num_of_f_un_rm_op_S FCVT_LU_S = ( (( 6 :: int)::ii))\ |\ num_of_f_un_rm_op_S FCVT_S_L = ( (( 7 :: int)::ii))\ |\ num_of_f_un_rm_op_S FCVT_S_LU = ( (( 8 :: int)::ii))\ \ \\val f_un_op_S_of_num : integer -> f_un_op_S\\ definition f_un_op_S_of_num :: \ int \ f_un_op_S \ where \ f_un_op_S_of_num arg1 = ( (let l__188 = arg1 in if (((l__188 = (( 0 :: int)::ii)))) then FCLASS_S else if (((l__188 = (( 1 :: int)::ii)))) then FMV_X_W else FMV_W_X))\ for arg1 :: " int " \ \\val num_of_f_un_op_S : f_un_op_S -> integer\\ fun num_of_f_un_op_S :: \ f_un_op_S \ int \ where \ num_of_f_un_op_S FCLASS_S = ( (( 0 :: int)::ii))\ |\ num_of_f_un_op_S FMV_X_W = ( (( 1 :: int)::ii))\ |\ num_of_f_un_op_S FMV_W_X = ( (( 2 :: int)::ii))\ \ \\val f_bin_op_S_of_num : integer -> f_bin_op_S\\ definition f_bin_op_S_of_num :: \ int \ f_bin_op_S \ where \ f_bin_op_S_of_num arg1 = ( (let l__181 = arg1 in if (((l__181 = (( 0 :: int)::ii)))) then FSGNJ_S else if (((l__181 = (( 1 :: int)::ii)))) then FSGNJN_S else if (((l__181 = (( 2 :: int)::ii)))) then FSGNJX_S else if (((l__181 = (( 3 :: int)::ii)))) then FMIN_S else if (((l__181 = (( 4 :: int)::ii)))) then FMAX_S else if (((l__181 = (( 5 :: int)::ii)))) then FEQ_S else if (((l__181 = (( 6 :: int)::ii)))) then FLT_S else FLE_S))\ for arg1 :: " int " \ \\val num_of_f_bin_op_S : f_bin_op_S -> integer\\ fun num_of_f_bin_op_S :: \ f_bin_op_S \ int \ where \ num_of_f_bin_op_S FSGNJ_S = ( (( 0 :: int)::ii))\ |\ num_of_f_bin_op_S FSGNJN_S = ( (( 1 :: int)::ii))\ |\ num_of_f_bin_op_S FSGNJX_S = ( (( 2 :: int)::ii))\ |\ num_of_f_bin_op_S FMIN_S = ( (( 3 :: int)::ii))\ |\ num_of_f_bin_op_S FMAX_S = ( (( 4 :: int)::ii))\ |\ num_of_f_bin_op_S FEQ_S = ( (( 5 :: int)::ii))\ |\ num_of_f_bin_op_S FLT_S = ( (( 6 :: int)::ii))\ |\ num_of_f_bin_op_S FLE_S = ( (( 7 :: int)::ii))\ \ \\val f_madd_op_D_of_num : integer -> f_madd_op_D\\ definition f_madd_op_D_of_num :: \ int \ f_madd_op_D \ where \ f_madd_op_D_of_num arg1 = ( (let l__178 = arg1 in if (((l__178 = (( 0 :: int)::ii)))) then FMADD_D else if (((l__178 = (( 1 :: int)::ii)))) then FMSUB_D else if (((l__178 = (( 2 :: int)::ii)))) then FNMSUB_D else FNMADD_D))\ for arg1 :: " int " \ \\val num_of_f_madd_op_D : f_madd_op_D -> integer\\ fun num_of_f_madd_op_D :: \ f_madd_op_D \ int \ where \ num_of_f_madd_op_D FMADD_D = ( (( 0 :: int)::ii))\ |\ num_of_f_madd_op_D FMSUB_D = ( (( 1 :: int)::ii))\ |\ num_of_f_madd_op_D FNMSUB_D = ( (( 2 :: int)::ii))\ |\ num_of_f_madd_op_D FNMADD_D = ( (( 3 :: int)::ii))\ \ \\val f_bin_rm_op_D_of_num : integer -> f_bin_rm_op_D\\ definition f_bin_rm_op_D_of_num :: \ int \ f_bin_rm_op_D \ where \ f_bin_rm_op_D_of_num arg1 = ( (let l__175 = arg1 in if (((l__175 = (( 0 :: int)::ii)))) then FADD_D else if (((l__175 = (( 1 :: int)::ii)))) then FSUB_D else if (((l__175 = (( 2 :: int)::ii)))) then FMUL_D else FDIV_D))\ for arg1 :: " int " \ \\val num_of_f_bin_rm_op_D : f_bin_rm_op_D -> integer\\ fun num_of_f_bin_rm_op_D :: \ f_bin_rm_op_D \ int \ where \ num_of_f_bin_rm_op_D FADD_D = ( (( 0 :: int)::ii))\ |\ num_of_f_bin_rm_op_D FSUB_D = ( (( 1 :: int)::ii))\ |\ num_of_f_bin_rm_op_D FMUL_D = ( (( 2 :: int)::ii))\ |\ num_of_f_bin_rm_op_D FDIV_D = ( (( 3 :: int)::ii))\ \ \\val f_un_rm_op_D_of_num : integer -> f_un_rm_op_D\\ definition f_un_rm_op_D_of_num :: \ int \ f_un_rm_op_D \ where \ f_un_rm_op_D_of_num arg1 = ( (let l__165 = arg1 in if (((l__165 = (( 0 :: int)::ii)))) then FSQRT_D else if (((l__165 = (( 1 :: int)::ii)))) then FCVT_W_D else if (((l__165 = (( 2 :: int)::ii)))) then FCVT_WU_D else if (((l__165 = (( 3 :: int)::ii)))) then FCVT_D_W else if (((l__165 = (( 4 :: int)::ii)))) then FCVT_D_WU else if (((l__165 = (( 5 :: int)::ii)))) then FCVT_S_D else if (((l__165 = (( 6 :: int)::ii)))) then FCVT_D_S else if (((l__165 = (( 7 :: int)::ii)))) then FCVT_L_D else if (((l__165 = (( 8 :: int)::ii)))) then FCVT_LU_D else if (((l__165 = (( 9 :: int)::ii)))) then FCVT_D_L else FCVT_D_LU))\ for arg1 :: " int " \ \\val num_of_f_un_rm_op_D : f_un_rm_op_D -> integer\\ fun num_of_f_un_rm_op_D :: \ f_un_rm_op_D \ int \ where \ num_of_f_un_rm_op_D FSQRT_D = ( (( 0 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_W_D = ( (( 1 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_WU_D = ( (( 2 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_D_W = ( (( 3 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_D_WU = ( (( 4 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_S_D = ( (( 5 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_D_S = ( (( 6 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_L_D = ( (( 7 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_LU_D = ( (( 8 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_D_L = ( (( 9 :: int)::ii))\ |\ num_of_f_un_rm_op_D FCVT_D_LU = ( (( 10 :: int)::ii))\ \ \\val f_bin_op_D_of_num : integer -> f_bin_op_D\\ definition f_bin_op_D_of_num :: \ int \ f_bin_op_D \ where \ f_bin_op_D_of_num arg1 = ( (let l__158 = arg1 in if (((l__158 = (( 0 :: int)::ii)))) then FSGNJ_D else if (((l__158 = (( 1 :: int)::ii)))) then FSGNJN_D else if (((l__158 = (( 2 :: int)::ii)))) then FSGNJX_D else if (((l__158 = (( 3 :: int)::ii)))) then FMIN_D else if (((l__158 = (( 4 :: int)::ii)))) then FMAX_D else if (((l__158 = (( 5 :: int)::ii)))) then FEQ_D else if (((l__158 = (( 6 :: int)::ii)))) then FLT_D else FLE_D))\ for arg1 :: " int " \ \\val num_of_f_bin_op_D : f_bin_op_D -> integer\\ fun num_of_f_bin_op_D :: \ f_bin_op_D \ int \ where \ num_of_f_bin_op_D FSGNJ_D = ( (( 0 :: int)::ii))\ |\ num_of_f_bin_op_D FSGNJN_D = ( (( 1 :: int)::ii))\ |\ num_of_f_bin_op_D FSGNJX_D = ( (( 2 :: int)::ii))\ |\ num_of_f_bin_op_D FMIN_D = ( (( 3 :: int)::ii))\ |\ num_of_f_bin_op_D FMAX_D = ( (( 4 :: int)::ii))\ |\ num_of_f_bin_op_D FEQ_D = ( (( 5 :: int)::ii))\ |\ num_of_f_bin_op_D FLT_D = ( (( 6 :: int)::ii))\ |\ num_of_f_bin_op_D FLE_D = ( (( 7 :: int)::ii))\ \ \\val f_un_op_D_of_num : integer -> f_un_op_D\\ definition f_un_op_D_of_num :: \ int \ f_un_op_D \ where \ f_un_op_D_of_num arg1 = ( (let l__156 = arg1 in if (((l__156 = (( 0 :: int)::ii)))) then FCLASS_D else if (((l__156 = (( 1 :: int)::ii)))) then FMV_X_D else FMV_D_X))\ for arg1 :: " int " \ \\val num_of_f_un_op_D : f_un_op_D -> integer\\ fun num_of_f_un_op_D :: \ f_un_op_D \ int \ where \ num_of_f_un_op_D FCLASS_D = ( (( 0 :: int)::ii))\ |\ num_of_f_un_op_D FMV_X_D = ( (( 1 :: int)::ii))\ |\ num_of_f_un_op_D FMV_D_X = ( (( 2 :: int)::ii))\ \ \\val rX : integer -> M (mword ty32)\\ definition rX :: \ int \((register_value),((32)Word.word),(exception))monad \ where \ rX r = ( (let l__124 = r in (if (((l__124 = (( 0 :: int)::ii)))) then return zero_reg else if (((l__124 = (( 1 :: int)::ii)))) then (read_reg x1_ref :: ( 32 Word.word) M) else if (((l__124 = (( 2 :: int)::ii)))) then (read_reg x2_ref :: ( 32 Word.word) M) else if (((l__124 = (( 3 :: int)::ii)))) then (read_reg x3_ref :: ( 32 Word.word) M) else if (((l__124 = (( 4 :: int)::ii)))) then (read_reg x4_ref :: ( 32 Word.word) M) else if (((l__124 = (( 5 :: int)::ii)))) then (read_reg x5_ref :: ( 32 Word.word) M) else if (((l__124 = (( 6 :: int)::ii)))) then (read_reg x6_ref :: ( 32 Word.word) M) else if (((l__124 = (( 7 :: int)::ii)))) then (read_reg x7_ref :: ( 32 Word.word) M) else if (((l__124 = (( 8 :: int)::ii)))) then (read_reg x8_ref :: ( 32 Word.word) M) else if (((l__124 = (( 9 :: int)::ii)))) then (read_reg x9_ref :: ( 32 Word.word) M) else if (((l__124 = (( 10 :: int)::ii)))) then (read_reg x10_ref :: ( 32 Word.word) M) else if (((l__124 = (( 11 :: int)::ii)))) then (read_reg x11_ref :: ( 32 Word.word) M) else if (((l__124 = (( 12 :: int)::ii)))) then (read_reg x12_ref :: ( 32 Word.word) M) else if (((l__124 = (( 13 :: int)::ii)))) then (read_reg x13_ref :: ( 32 Word.word) M) else if (((l__124 = (( 14 :: int)::ii)))) then (read_reg x14_ref :: ( 32 Word.word) M) else if (((l__124 = (( 15 :: int)::ii)))) then (read_reg x15_ref :: ( 32 Word.word) M) else if (((l__124 = (( 16 :: int)::ii)))) then (read_reg x16_ref :: ( 32 Word.word) M) else if (((l__124 = (( 17 :: int)::ii)))) then (read_reg x17_ref :: ( 32 Word.word) M) else if (((l__124 = (( 18 :: int)::ii)))) then (read_reg x18_ref :: ( 32 Word.word) M) else if (((l__124 = (( 19 :: int)::ii)))) then (read_reg x19_ref :: ( 32 Word.word) M) else if (((l__124 = (( 20 :: int)::ii)))) then (read_reg x20_ref :: ( 32 Word.word) M) else if (((l__124 = (( 21 :: int)::ii)))) then (read_reg x21_ref :: ( 32 Word.word) M) else if (((l__124 = (( 22 :: int)::ii)))) then (read_reg x22_ref :: ( 32 Word.word) M) else if (((l__124 = (( 23 :: int)::ii)))) then (read_reg x23_ref :: ( 32 Word.word) M) else if (((l__124 = (( 24 :: int)::ii)))) then (read_reg x24_ref :: ( 32 Word.word) M) else if (((l__124 = (( 25 :: int)::ii)))) then (read_reg x25_ref :: ( 32 Word.word) M) else if (((l__124 = (( 26 :: int)::ii)))) then (read_reg x26_ref :: ( 32 Word.word) M) else if (((l__124 = (( 27 :: int)::ii)))) then (read_reg x27_ref :: ( 32 Word.word) M) else if (((l__124 = (( 28 :: int)::ii)))) then (read_reg x28_ref :: ( 32 Word.word) M) else if (((l__124 = (( 29 :: int)::ii)))) then (read_reg x29_ref :: ( 32 Word.word) M) else if (((l__124 = (( 30 :: int)::ii)))) then (read_reg x30_ref :: ( 32 Word.word) M) else if (((l__124 = (( 31 :: int)::ii)))) then (read_reg x31_ref :: ( 32 Word.word) M) else assert_exp False (''invalid register number'') \ exit0 () ) \ ((\ (v :: regtype) . return ((regval_from_reg v :: 32 Word.word))))))\ for r :: " int " \ \\val rvfi_wX : integer -> mword ty32 -> unit\\ definition rvfi_wX :: \ int \(32)Word.word \ unit \ where \ rvfi_wX r v = ( () )\ for r :: " int " and v :: "(32)Word.word " \ \\val wX : integer -> mword ty32 -> M unit\\ definition wX :: \ int \(32)Word.word \((register_value),(unit),(exception))monad \ where \ wX r in_v = ( (let v = ((regval_into_reg in_v :: 32 Word.word)) in (let l__92 = r in (if (((l__92 = (( 0 :: int)::ii)))) then return () else if (((l__92 = (( 1 :: int)::ii)))) then write_reg x1_ref v else if (((l__92 = (( 2 :: int)::ii)))) then write_reg x2_ref v else if (((l__92 = (( 3 :: int)::ii)))) then write_reg x3_ref v else if (((l__92 = (( 4 :: int)::ii)))) then write_reg x4_ref v else if (((l__92 = (( 5 :: int)::ii)))) then write_reg x5_ref v else if (((l__92 = (( 6 :: int)::ii)))) then write_reg x6_ref v else if (((l__92 = (( 7 :: int)::ii)))) then write_reg x7_ref v else if (((l__92 = (( 8 :: int)::ii)))) then write_reg x8_ref v else if (((l__92 = (( 9 :: int)::ii)))) then write_reg x9_ref v else if (((l__92 = (( 10 :: int)::ii)))) then write_reg x10_ref v else if (((l__92 = (( 11 :: int)::ii)))) then write_reg x11_ref v else if (((l__92 = (( 12 :: int)::ii)))) then write_reg x12_ref v else if (((l__92 = (( 13 :: int)::ii)))) then write_reg x13_ref v else if (((l__92 = (( 14 :: int)::ii)))) then write_reg x14_ref v else if (((l__92 = (( 15 :: int)::ii)))) then write_reg x15_ref v else if (((l__92 = (( 16 :: int)::ii)))) then write_reg x16_ref v else if (((l__92 = (( 17 :: int)::ii)))) then write_reg x17_ref v else if (((l__92 = (( 18 :: int)::ii)))) then write_reg x18_ref v else if (((l__92 = (( 19 :: int)::ii)))) then write_reg x19_ref v else if (((l__92 = (( 20 :: int)::ii)))) then write_reg x20_ref v else if (((l__92 = (( 21 :: int)::ii)))) then write_reg x21_ref v else if (((l__92 = (( 22 :: int)::ii)))) then write_reg x22_ref v else if (((l__92 = (( 23 :: int)::ii)))) then write_reg x23_ref v else if (((l__92 = (( 24 :: int)::ii)))) then write_reg x24_ref v else if (((l__92 = (( 25 :: int)::ii)))) then write_reg x25_ref v else if (((l__92 = (( 26 :: int)::ii)))) then write_reg x26_ref v else if (((l__92 = (( 27 :: int)::ii)))) then write_reg x27_ref v else if (((l__92 = (( 28 :: int)::ii)))) then write_reg x28_ref v else if (((l__92 = (( 29 :: int)::ii)))) then write_reg x29_ref v else if (((l__92 = (( 30 :: int)::ii)))) then write_reg x30_ref v else if (((l__92 = (( 31 :: int)::ii)))) then write_reg x31_ref v else assert_exp False (''invalid register number'') \ exit0 () ) \ return (if (((r \ (( 0 :: int)::ii)))) then (let (_ :: unit) = (rvfi_wX r in_v) in if ((get_config_print_reg () )) then print_dbg (((@) (''x'') (((@) ((stringFromInteger r)) (((@) ('' <- '') ((RegStr v)))))))) else () ) else () ))))\ for r :: " int " and in_v :: "(32)Word.word " \ \\val rX_bits : mword ty5 -> M (mword ty32)\\ definition rX_bits :: \(5)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ rX_bits i = ( (rX ((Word.uint i)) :: ( 32 Word.word) M))\ for i :: "(5)Word.word " \ \\val wX_bits : mword ty5 -> mword ty32 -> M unit\\ definition wX_bits :: \(5)Word.word \(32)Word.word \((register_value),(unit),(exception))monad \ where \ wX_bits (i :: 5 bits) (data :: xlenbits) = ( wX ((Word.uint i)) data )\ for i :: "(5)Word.word " and data :: "(32)Word.word " definition reg_name_abi :: \(5)Word.word \((register_value),(string),(exception))monad \ where \ reg_name_abi r = ( (let b__0 = r in if (((b__0 = ( 0b00000 :: 5 Word.word)))) then return (''zero'') else if (((b__0 = ( 0b00001 :: 5 Word.word)))) then return (''ra'') else if (((b__0 = ( 0b00010 :: 5 Word.word)))) then return (''sp'') else if (((b__0 = ( 0b00011 :: 5 Word.word)))) then return (''gp'') else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then return (''tp'') else if (((b__0 = ( 0b00101 :: 5 Word.word)))) then return (''t0'') else if (((b__0 = ( 0b00110 :: 5 Word.word)))) then return (''t1'') else if (((b__0 = ( 0b00111 :: 5 Word.word)))) then return (''t2'') else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then return (''fp'') else if (((b__0 = ( 0b01001 :: 5 Word.word)))) then return (''s1'') else if (((b__0 = ( 0b01010 :: 5 Word.word)))) then return (''a0'') else if (((b__0 = ( 0b01011 :: 5 Word.word)))) then return (''a1'') else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then return (''a2'') else if (((b__0 = ( 0b01101 :: 5 Word.word)))) then return (''a3'') else if (((b__0 = ( 0b01110 :: 5 Word.word)))) then return (''a4'') else if (((b__0 = ( 0b01111 :: 5 Word.word)))) then return (''a5'') else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then return (''a6'') else if (((b__0 = ( 0b10001 :: 5 Word.word)))) then return (''a7'') else if (((b__0 = ( 0b10010 :: 5 Word.word)))) then return (''s2'') else if (((b__0 = ( 0b10011 :: 5 Word.word)))) then return (''s3'') else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then return (''s4'') else if (((b__0 = ( 0b10101 :: 5 Word.word)))) then return (''s5'') else if (((b__0 = ( 0b10110 :: 5 Word.word)))) then return (''s6'') else if (((b__0 = ( 0b10111 :: 5 Word.word)))) then return (''s7'') else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then return (''s8'') else if (((b__0 = ( 0b11001 :: 5 Word.word)))) then return (''s9'') else if (((b__0 = ( 0b11010 :: 5 Word.word)))) then return (''s10'') else if (((b__0 = ( 0b11011 :: 5 Word.word)))) then return (''s11'') else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then return (''t3'') else if (((b__0 = ( 0b11101 :: 5 Word.word)))) then return (''t4'') else if (((b__0 = ( 0b11110 :: 5 Word.word)))) then return (''t5'') else if (((b__0 = ( 0b11111 :: 5 Word.word)))) then return (''t6'') else assert_exp False (''Pattern match failure at model/riscv_regs.sail 154:2 - 187:3'') \ exit0 () ))\ for r :: "(5)Word.word " \ \\val reg_name_forwards : mword ty5 -> M string\\ \ \\val reg_name_backwards : string -> M (mword ty5)\\ \ \\val reg_name_forwards_matches : mword ty5 -> bool\\ \ \\val reg_name_backwards_matches : string -> bool\\ \ \\val reg_name_matches_prefix : string -> maybe ((mword ty5 * ii))\\ definition reg_name_forwards :: \(5)Word.word \((register_value),(string),(exception))monad \ where \ reg_name_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00000 :: 5 Word.word)))) then return (''zero'') else if (((b__0 = ( 0b00001 :: 5 Word.word)))) then return (''ra'') else if (((b__0 = ( 0b00010 :: 5 Word.word)))) then return (''sp'') else if (((b__0 = ( 0b00011 :: 5 Word.word)))) then return (''gp'') else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then return (''tp'') else if (((b__0 = ( 0b00101 :: 5 Word.word)))) then return (''t0'') else if (((b__0 = ( 0b00110 :: 5 Word.word)))) then return (''t1'') else if (((b__0 = ( 0b00111 :: 5 Word.word)))) then return (''t2'') else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then return (''fp'') else if (((b__0 = ( 0b01001 :: 5 Word.word)))) then return (''s1'') else if (((b__0 = ( 0b01010 :: 5 Word.word)))) then return (''a0'') else if (((b__0 = ( 0b01011 :: 5 Word.word)))) then return (''a1'') else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then return (''a2'') else if (((b__0 = ( 0b01101 :: 5 Word.word)))) then return (''a3'') else if (((b__0 = ( 0b01110 :: 5 Word.word)))) then return (''a4'') else if (((b__0 = ( 0b01111 :: 5 Word.word)))) then return (''a5'') else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then return (''a6'') else if (((b__0 = ( 0b10001 :: 5 Word.word)))) then return (''a7'') else if (((b__0 = ( 0b10010 :: 5 Word.word)))) then return (''s2'') else if (((b__0 = ( 0b10011 :: 5 Word.word)))) then return (''s3'') else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then return (''s4'') else if (((b__0 = ( 0b10101 :: 5 Word.word)))) then return (''s5'') else if (((b__0 = ( 0b10110 :: 5 Word.word)))) then return (''s6'') else if (((b__0 = ( 0b10111 :: 5 Word.word)))) then return (''s7'') else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then return (''s8'') else if (((b__0 = ( 0b11001 :: 5 Word.word)))) then return (''s9'') else if (((b__0 = ( 0b11010 :: 5 Word.word)))) then return (''s10'') else if (((b__0 = ( 0b11011 :: 5 Word.word)))) then return (''s11'') else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then return (''t3'') else if (((b__0 = ( 0b11101 :: 5 Word.word)))) then return (''t4'') else if (((b__0 = ( 0b11110 :: 5 Word.word)))) then return (''t5'') else if (((b__0 = ( 0b11111 :: 5 Word.word)))) then return (''t6'') else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(5)Word.word " definition reg_name_backwards :: \ string \((register_value),((5)Word.word),(exception))monad \ where \ reg_name_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''zero'')))) then return ( 0b00000 :: 5 Word.word) else if (((p00 = (''ra'')))) then return ( 0b00001 :: 5 Word.word) else if (((p00 = (''sp'')))) then return ( 0b00010 :: 5 Word.word) else if (((p00 = (''gp'')))) then return ( 0b00011 :: 5 Word.word) else if (((p00 = (''tp'')))) then return ( 0b00100 :: 5 Word.word) else if (((p00 = (''t0'')))) then return ( 0b00101 :: 5 Word.word) else if (((p00 = (''t1'')))) then return ( 0b00110 :: 5 Word.word) else if (((p00 = (''t2'')))) then return ( 0b00111 :: 5 Word.word) else if (((p00 = (''fp'')))) then return ( 0b01000 :: 5 Word.word) else if (((p00 = (''s1'')))) then return ( 0b01001 :: 5 Word.word) else if (((p00 = (''a0'')))) then return ( 0b01010 :: 5 Word.word) else if (((p00 = (''a1'')))) then return ( 0b01011 :: 5 Word.word) else if (((p00 = (''a2'')))) then return ( 0b01100 :: 5 Word.word) else if (((p00 = (''a3'')))) then return ( 0b01101 :: 5 Word.word) else if (((p00 = (''a4'')))) then return ( 0b01110 :: 5 Word.word) else if (((p00 = (''a5'')))) then return ( 0b01111 :: 5 Word.word) else if (((p00 = (''a6'')))) then return ( 0b10000 :: 5 Word.word) else if (((p00 = (''a7'')))) then return ( 0b10001 :: 5 Word.word) else if (((p00 = (''s2'')))) then return ( 0b10010 :: 5 Word.word) else if (((p00 = (''s3'')))) then return ( 0b10011 :: 5 Word.word) else if (((p00 = (''s4'')))) then return ( 0b10100 :: 5 Word.word) else if (((p00 = (''s5'')))) then return ( 0b10101 :: 5 Word.word) else if (((p00 = (''s6'')))) then return ( 0b10110 :: 5 Word.word) else if (((p00 = (''s7'')))) then return ( 0b10111 :: 5 Word.word) else if (((p00 = (''s8'')))) then return ( 0b11000 :: 5 Word.word) else if (((p00 = (''s9'')))) then return ( 0b11001 :: 5 Word.word) else if (((p00 = (''s10'')))) then return ( 0b11010 :: 5 Word.word) else if (((p00 = (''s11'')))) then return ( 0b11011 :: 5 Word.word) else if (((p00 = (''t3'')))) then return ( 0b11100 :: 5 Word.word) else if (((p00 = (''t4'')))) then return ( 0b11101 :: 5 Word.word) else if (((p00 = (''t5'')))) then return ( 0b11110 :: 5 Word.word) else if (((p00 = (''t6'')))) then return ( 0b11111 :: 5 Word.word) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition reg_name_forwards_matches :: \(5)Word.word \ bool \ where \ reg_name_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11111 :: 5 Word.word)))) then True else False))\ for arg1 :: "(5)Word.word " definition reg_name_backwards_matches :: \ string \ bool \ where \ reg_name_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''zero'')))) then True else if (((p00 = (''ra'')))) then True else if (((p00 = (''sp'')))) then True else if (((p00 = (''gp'')))) then True else if (((p00 = (''tp'')))) then True else if (((p00 = (''t0'')))) then True else if (((p00 = (''t1'')))) then True else if (((p00 = (''t2'')))) then True else if (((p00 = (''fp'')))) then True else if (((p00 = (''s1'')))) then True else if (((p00 = (''a0'')))) then True else if (((p00 = (''a1'')))) then True else if (((p00 = (''a2'')))) then True else if (((p00 = (''a3'')))) then True else if (((p00 = (''a4'')))) then True else if (((p00 = (''a5'')))) then True else if (((p00 = (''a6'')))) then True else if (((p00 = (''a7'')))) then True else if (((p00 = (''s2'')))) then True else if (((p00 = (''s3'')))) then True else if (((p00 = (''s4'')))) then True else if (((p00 = (''s5'')))) then True else if (((p00 = (''s6'')))) then True else if (((p00 = (''s7'')))) then True else if (((p00 = (''s8'')))) then True else if (((p00 = (''s9'')))) then True else if (((p00 = (''s10'')))) then True else if (((p00 = (''s11'')))) then True else if (((p00 = (''t3'')))) then True else if (((p00 = (''t4'')))) then True else if (((p00 = (''t5'')))) then True else if (((p00 = (''t6'')))) then True else False))\ for arg1 :: " string " \ \\val _s164_ : string -> maybe string\\ definition s164 :: \ string \(string)option \ where \ s164 s1650 = ( (let s1660 = s1650 in if ((string_startswith s1660 (''t6''))) then (case ((string_drop s1660 ((string_length (''t6''))))) of s1 => Some s1 ) else None))\ for s1650 :: " string " \ \\val _s160_ : string -> maybe string\\ definition s160 :: \ string \(string)option \ where \ s160 s1611 = ( (let s1620 = s1611 in if ((string_startswith s1620 (''t5''))) then (case ((string_drop s1620 ((string_length (''t5''))))) of s1 => Some s1 ) else None))\ for s1611 :: " string " \ \\val _s156_ : string -> maybe string\\ definition s156 :: \ string \(string)option \ where \ s156 s1570 = ( (let s1580 = s1570 in if ((string_startswith s1580 (''t4''))) then (case ((string_drop s1580 ((string_length (''t4''))))) of s1 => Some s1 ) else None))\ for s1570 :: " string " \ \\val _s152_ : string -> maybe string\\ definition s152 :: \ string \(string)option \ where \ s152 s1530 = ( (let s1540 = s1530 in if ((string_startswith s1540 (''t3''))) then (case ((string_drop s1540 ((string_length (''t3''))))) of s1 => Some s1 ) else None))\ for s1530 :: " string " \ \\val _s148_ : string -> maybe string\\ definition s148 :: \ string \(string)option \ where \ s148 s1490 = ( (let s1500 = s1490 in if ((string_startswith s1500 (''s11''))) then (case ((string_drop s1500 ((string_length (''s11''))))) of s1 => Some s1 ) else None))\ for s1490 :: " string " \ \\val _s144_ : string -> maybe string\\ definition s144 :: \ string \(string)option \ where \ s144 s1450 = ( (let s1461 = s1450 in if ((string_startswith s1461 (''s10''))) then (case ((string_drop s1461 ((string_length (''s10''))))) of s1 => Some s1 ) else None))\ for s1450 :: " string " \ \\val _s140_ : string -> maybe string\\ definition s140 :: \ string \(string)option \ where \ s140 s1410 = ( (let s1420 = s1410 in if ((string_startswith s1420 (''s9''))) then (case ((string_drop s1420 ((string_length (''s9''))))) of s1 => Some s1 ) else None))\ for s1410 :: " string " \ \\val _s136_ : string -> maybe string\\ definition s136 :: \ string \(string)option \ where \ s136 s1370 = ( (let s1380 = s1370 in if ((string_startswith s1380 (''s8''))) then (case ((string_drop s1380 ((string_length (''s8''))))) of s1 => Some s1 ) else None))\ for s1370 :: " string " \ \\val _s132_ : string -> maybe string\\ definition s132 :: \ string \(string)option \ where \ s132 s1330 = ( (let s1340 = s1330 in if ((string_startswith s1340 (''s7''))) then (case ((string_drop s1340 ((string_length (''s7''))))) of s1 => Some s1 ) else None))\ for s1330 :: " string " \ \\val _s128_ : string -> maybe string\\ definition s128 :: \ string \(string)option \ where \ s128 s1290 = ( (let s1300 = s1290 in if ((string_startswith s1300 (''s6''))) then (case ((string_drop s1300 ((string_length (''s6''))))) of s1 => Some s1 ) else None))\ for s1290 :: " string " \ \\val _s124_ : string -> maybe string\\ definition s124 :: \ string \(string)option \ where \ s124 s1250 = ( (let s1260 = s1250 in if ((string_startswith s1260 (''s5''))) then (case ((string_drop s1260 ((string_length (''s5''))))) of s1 => Some s1 ) else None))\ for s1250 :: " string " \ \\val _s120_ : string -> maybe string\\ definition s120 :: \ string \(string)option \ where \ s120 s1211 = ( (let s1220 = s1211 in if ((string_startswith s1220 (''s4''))) then (case ((string_drop s1220 ((string_length (''s4''))))) of s1 => Some s1 ) else None))\ for s1211 :: " string " \ \\val _s116_ : string -> maybe string\\ definition s116 :: \ string \(string)option \ where \ s116 s1171 = ( (let s1180 = s1171 in if ((string_startswith s1180 (''s3''))) then (case ((string_drop s1180 ((string_length (''s3''))))) of s1 => Some s1 ) else None))\ for s1171 :: " string " \ \\val _s112_ : string -> maybe string\\ definition s112 :: \ string \(string)option \ where \ s112 s1131 = ( (let s1140 = s1131 in if ((string_startswith s1140 (''s2''))) then (case ((string_drop s1140 ((string_length (''s2''))))) of s1 => Some s1 ) else None))\ for s1131 :: " string " \ \\val _s108_ : string -> maybe string\\ definition s108 :: \ string \(string)option \ where \ s108 s1091 = ( (let s1100 = s1091 in if ((string_startswith s1100 (''a7''))) then (case ((string_drop s1100 ((string_length (''a7''))))) of s1 => Some s1 ) else None))\ for s1091 :: " string " \ \\val _s104_ : string -> maybe string\\ definition s104 :: \ string \(string)option \ where \ s104 s1051 = ( (let s1060 = s1051 in if ((string_startswith s1060 (''a6''))) then (case ((string_drop s1060 ((string_length (''a6''))))) of s1 => Some s1 ) else None))\ for s1051 :: " string " \ \\val _s100_ : string -> maybe string\\ definition s100 :: \ string \(string)option \ where \ s100 s1011 = ( (let s1020 = s1011 in if ((string_startswith s1020 (''a5''))) then (case ((string_drop s1020 ((string_length (''a5''))))) of s1 => Some s1 ) else None))\ for s1011 :: " string " \ \\val _s96_ : string -> maybe string\\ definition s96 :: \ string \(string)option \ where \ s96 s970 = ( (let s980 = s970 in if ((string_startswith s980 (''a4''))) then (case ((string_drop s980 ((string_length (''a4''))))) of s1 => Some s1 ) else None))\ for s970 :: " string " \ \\val _s92_ : string -> maybe string\\ definition s92 :: \ string \(string)option \ where \ s92 s930 = ( (let s940 = s930 in if ((string_startswith s940 (''a3''))) then (case ((string_drop s940 ((string_length (''a3''))))) of s1 => Some s1 ) else None))\ for s930 :: " string " \ \\val _s88_ : string -> maybe string\\ definition s88 :: \ string \(string)option \ where \ s88 s890 = ( (let s900 = s890 in if ((string_startswith s900 (''a2''))) then (case ((string_drop s900 ((string_length (''a2''))))) of s1 => Some s1 ) else None))\ for s890 :: " string " \ \\val _s84_ : string -> maybe string\\ definition s84 :: \ string \(string)option \ where \ s84 s850 = ( (let s860 = s850 in if ((string_startswith s860 (''a1''))) then (case ((string_drop s860 ((string_length (''a1''))))) of s1 => Some s1 ) else None))\ for s850 :: " string " \ \\val _s80_ : string -> maybe string\\ definition s80 :: \ string \(string)option \ where \ s80 s810 = ( (let s820 = s810 in if ((string_startswith s820 (''a0''))) then (case ((string_drop s820 ((string_length (''a0''))))) of s1 => Some s1 ) else None))\ for s810 :: " string " \ \\val _s76_ : string -> maybe string\\ definition s76 :: \ string \(string)option \ where \ s76 s770 = ( (let s780 = s770 in if ((string_startswith s780 (''s1''))) then (case ((string_drop s780 ((string_length (''s1''))))) of s1 => Some s1 ) else None))\ for s770 :: " string " \ \\val _s72_ : string -> maybe string\\ definition s72 :: \ string \(string)option \ where \ s72 s730 = ( (let s741 = s730 in if ((string_startswith s741 (''fp''))) then (case ((string_drop s741 ((string_length (''fp''))))) of s1 => Some s1 ) else None))\ for s730 :: " string " \ \\val _s68_ : string -> maybe string\\ definition s68 :: \ string \(string)option \ where \ s68 s690 = ( (let s701 = s690 in if ((string_startswith s701 (''t2''))) then (case ((string_drop s701 ((string_length (''t2''))))) of s1 => Some s1 ) else None))\ for s690 :: " string " \ \\val _s64_ : string -> maybe string\\ definition s64 :: \ string \(string)option \ where \ s64 s650 = ( (let s661 = s650 in if ((string_startswith s661 (''t1''))) then (case ((string_drop s661 ((string_length (''t1''))))) of s1 => Some s1 ) else None))\ for s650 :: " string " \ \\val _s60_ : string -> maybe string\\ definition s60 :: \ string \(string)option \ where \ s60 s610 = ( (let s621 = s610 in if ((string_startswith s621 (''t0''))) then (case ((string_drop s621 ((string_length (''t0''))))) of s1 => Some s1 ) else None))\ for s610 :: " string " \ \\val _s56_ : string -> maybe string\\ definition s56 :: \ string \(string)option \ where \ s56 s570 = ( (let s581 = s570 in if ((string_startswith s581 (''tp''))) then (case ((string_drop s581 ((string_length (''tp''))))) of s1 => Some s1 ) else None))\ for s570 :: " string " \ \\val _s52_ : string -> maybe string\\ definition s52 :: \ string \(string)option \ where \ s52 s530 = ( (let s541 = s530 in if ((string_startswith s541 (''gp''))) then (case ((string_drop s541 ((string_length (''gp''))))) of s1 => Some s1 ) else None))\ for s530 :: " string " \ \\val _s48_ : string -> maybe string\\ definition s48 :: \ string \(string)option \ where \ s48 s490 = ( (let s501 = s490 in if ((string_startswith s501 (''sp''))) then (case ((string_drop s501 ((string_length (''sp''))))) of s1 => Some s1 ) else None))\ for s490 :: " string " \ \\val _s44_ : string -> maybe string\\ definition s44 :: \ string \(string)option \ where \ s44 s450 = ( (let s461 = s450 in if ((string_startswith s461 (''ra''))) then (case ((string_drop s461 ((string_length (''ra''))))) of s1 => Some s1 ) else None))\ for s450 :: " string " \ \\val _s40_ : string -> maybe string\\ definition s40 :: \ string \(string)option \ where \ s40 s410 = ( (let s421 = s410 in if ((string_startswith s421 (''zero''))) then (case ((string_drop s421 ((string_length (''zero''))))) of s1 => Some s1 ) else None))\ for s410 :: " string " definition reg_name_matches_prefix :: \ string \((5)Word.word*int)option \ where \ reg_name_matches_prefix arg1 = ( (let s430 = arg1 in if ((case ((s40 s430)) of Some (s1) => True | _ => False )) then (case s40 s430 of (Some (s1)) => Some (( 0b00000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s44 s430)) of Some (s1) => True | _ => False )) then (case s44 s430 of (Some (s1)) => Some (( 0b00001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s48 s430)) of Some (s1) => True | _ => False )) then (case s48 s430 of (Some (s1)) => Some (( 0b00010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s52 s430)) of Some (s1) => True | _ => False )) then (case s52 s430 of (Some (s1)) => Some (( 0b00011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s56 s430)) of Some (s1) => True | _ => False )) then (case s56 s430 of (Some (s1)) => Some (( 0b00100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s60 s430)) of Some (s1) => True | _ => False )) then (case s60 s430 of (Some (s1)) => Some (( 0b00101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s64 s430)) of Some (s1) => True | _ => False )) then (case s64 s430 of (Some (s1)) => Some (( 0b00110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s68 s430)) of Some (s1) => True | _ => False )) then (case s68 s430 of (Some (s1)) => Some (( 0b00111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s72 s430)) of Some (s1) => True | _ => False )) then (case s72 s430 of (Some (s1)) => Some (( 0b01000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s76 s430)) of Some (s1) => True | _ => False )) then (case s76 s430 of (Some (s1)) => Some (( 0b01001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s80 s430)) of Some (s1) => True | _ => False )) then (case s80 s430 of (Some (s1)) => Some (( 0b01010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s84 s430)) of Some (s1) => True | _ => False )) then (case s84 s430 of (Some (s1)) => Some (( 0b01011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s88 s430)) of Some (s1) => True | _ => False )) then (case s88 s430 of (Some (s1)) => Some (( 0b01100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s92 s430)) of Some (s1) => True | _ => False )) then (case s92 s430 of (Some (s1)) => Some (( 0b01101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s96 s430)) of Some (s1) => True | _ => False )) then (case s96 s430 of (Some (s1)) => Some (( 0b01110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s100 s430)) of Some (s1) => True | _ => False )) then (case s100 s430 of (Some (s1)) => Some (( 0b01111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s104 s430)) of Some (s1) => True | _ => False )) then (case s104 s430 of (Some (s1)) => Some (( 0b10000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s108 s430)) of Some (s1) => True | _ => False )) then (case s108 s430 of (Some (s1)) => Some (( 0b10001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s112 s430)) of Some (s1) => True | _ => False )) then (case s112 s430 of (Some (s1)) => Some (( 0b10010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s116 s430)) of Some (s1) => True | _ => False )) then (case s116 s430 of (Some (s1)) => Some (( 0b10011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s120 s430)) of Some (s1) => True | _ => False )) then (case s120 s430 of (Some (s1)) => Some (( 0b10100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s124 s430)) of Some (s1) => True | _ => False )) then (case s124 s430 of (Some (s1)) => Some (( 0b10101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s128 s430)) of Some (s1) => True | _ => False )) then (case s128 s430 of (Some (s1)) => Some (( 0b10110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s132 s430)) of Some (s1) => True | _ => False )) then (case s132 s430 of (Some (s1)) => Some (( 0b10111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s136 s430)) of Some (s1) => True | _ => False )) then (case s136 s430 of (Some (s1)) => Some (( 0b11000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s140 s430)) of Some (s1) => True | _ => False )) then (case s140 s430 of (Some (s1)) => Some (( 0b11001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s144 s430)) of Some (s1) => True | _ => False )) then (case s144 s430 of (Some (s1)) => Some (( 0b11010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s148 s430)) of Some (s1) => True | _ => False )) then (case s148 s430 of (Some (s1)) => Some (( 0b11011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s152 s430)) of Some (s1) => True | _ => False )) then (case s152 s430 of (Some (s1)) => Some (( 0b11100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s156 s430)) of Some (s1) => True | _ => False )) then (case s156 s430 of (Some (s1)) => Some (( 0b11101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s160 s430)) of Some (s1) => True | _ => False )) then (case s160 s430 of (Some (s1)) => Some (( 0b11110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s164 s430)) of Some (s1) => True | _ => False )) then (case s164 s430 of (Some (s1)) => Some (( 0b11111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val creg_name_forwards : mword ty3 -> M string\\ \ \\val creg_name_backwards : string -> M (mword ty3)\\ \ \\val creg_name_forwards_matches : mword ty3 -> bool\\ \ \\val creg_name_backwards_matches : string -> bool\\ \ \\val creg_name_matches_prefix : string -> maybe ((mword ty3 * ii))\\ definition creg_name_forwards :: \(3)Word.word \((register_value),(string),(exception))monad \ where \ creg_name_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then return (''s0'') else if (((b__0 = ( 0b001 :: 3 Word.word)))) then return (''s1'') else if (((b__0 = ( 0b010 :: 3 Word.word)))) then return (''a0'') else if (((b__0 = ( 0b011 :: 3 Word.word)))) then return (''a1'') else if (((b__0 = ( 0b100 :: 3 Word.word)))) then return (''a2'') else if (((b__0 = ( 0b101 :: 3 Word.word)))) then return (''a3'') else if (((b__0 = ( 0b110 :: 3 Word.word)))) then return (''a4'') else if (((b__0 = ( 0b111 :: 3 Word.word)))) then return (''a5'') else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(3)Word.word " definition creg_name_backwards :: \ string \((register_value),((3)Word.word),(exception))monad \ where \ creg_name_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''s0'')))) then return ( 0b000 :: 3 Word.word) else if (((p00 = (''s1'')))) then return ( 0b001 :: 3 Word.word) else if (((p00 = (''a0'')))) then return ( 0b010 :: 3 Word.word) else if (((p00 = (''a1'')))) then return ( 0b011 :: 3 Word.word) else if (((p00 = (''a2'')))) then return ( 0b100 :: 3 Word.word) else if (((p00 = (''a3'')))) then return ( 0b101 :: 3 Word.word) else if (((p00 = (''a4'')))) then return ( 0b110 :: 3 Word.word) else if (((p00 = (''a5'')))) then return ( 0b111 :: 3 Word.word) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition creg_name_forwards_matches :: \(3)Word.word \ bool \ where \ creg_name_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then True else if (((b__0 = ( 0b001 :: 3 Word.word)))) then True else if (((b__0 = ( 0b010 :: 3 Word.word)))) then True else if (((b__0 = ( 0b011 :: 3 Word.word)))) then True else if (((b__0 = ( 0b100 :: 3 Word.word)))) then True else if (((b__0 = ( 0b101 :: 3 Word.word)))) then True else if (((b__0 = ( 0b110 :: 3 Word.word)))) then True else if (((b__0 = ( 0b111 :: 3 Word.word)))) then True else False))\ for arg1 :: "(3)Word.word " definition creg_name_backwards_matches :: \ string \ bool \ where \ creg_name_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''s0'')))) then True else if (((p00 = (''s1'')))) then True else if (((p00 = (''a0'')))) then True else if (((p00 = (''a1'')))) then True else if (((p00 = (''a2'')))) then True else if (((p00 = (''a3'')))) then True else if (((p00 = (''a4'')))) then True else if (((p00 = (''a5'')))) then True else False))\ for arg1 :: " string " \ \\val _s196_ : string -> maybe string\\ definition s196 :: \ string \(string)option \ where \ s196 s1970 = ( (let s1980 = s1970 in if ((string_startswith s1980 (''a5''))) then (case ((string_drop s1980 ((string_length (''a5''))))) of s1 => Some s1 ) else None))\ for s1970 :: " string " \ \\val _s192_ : string -> maybe string\\ definition s192 :: \ string \(string)option \ where \ s192 s1931 = ( (let s1940 = s1931 in if ((string_startswith s1940 (''a4''))) then (case ((string_drop s1940 ((string_length (''a4''))))) of s1 => Some s1 ) else None))\ for s1931 :: " string " \ \\val _s188_ : string -> maybe string\\ definition s188 :: \ string \(string)option \ where \ s188 s1891 = ( (let s1900 = s1891 in if ((string_startswith s1900 (''a3''))) then (case ((string_drop s1900 ((string_length (''a3''))))) of s1 => Some s1 ) else None))\ for s1891 :: " string " \ \\val _s184_ : string -> maybe string\\ definition s184 :: \ string \(string)option \ where \ s184 s1850 = ( (let s1860 = s1850 in if ((string_startswith s1860 (''a2''))) then (case ((string_drop s1860 ((string_length (''a2''))))) of s1 => Some s1 ) else None))\ for s1850 :: " string " \ \\val _s180_ : string -> maybe string\\ definition s180 :: \ string \(string)option \ where \ s180 s1810 = ( (let s1820 = s1810 in if ((string_startswith s1820 (''a1''))) then (case ((string_drop s1820 ((string_length (''a1''))))) of s1 => Some s1 ) else None))\ for s1810 :: " string " \ \\val _s176_ : string -> maybe string\\ definition s176 :: \ string \(string)option \ where \ s176 s1770 = ( (let s1780 = s1770 in if ((string_startswith s1780 (''a0''))) then (case ((string_drop s1780 ((string_length (''a0''))))) of s1 => Some s1 ) else None))\ for s1770 :: " string " \ \\val _s172_ : string -> maybe string\\ definition s172 :: \ string \(string)option \ where \ s172 s1730 = ( (let s1740 = s1730 in if ((string_startswith s1740 (''s1''))) then (case ((string_drop s1740 ((string_length (''s1''))))) of s1 => Some s1 ) else None))\ for s1730 :: " string " \ \\val _s168_ : string -> maybe string\\ definition s168 :: \ string \(string)option \ where \ s168 s1691 = ( (let s1700 = s1691 in if ((string_startswith s1700 (''s0''))) then (case ((string_drop s1700 ((string_length (''s0''))))) of s1 => Some s1 ) else None))\ for s1691 :: " string " definition creg_name_matches_prefix :: \ string \((3)Word.word*int)option \ where \ creg_name_matches_prefix arg1 = ( (let s1710 = arg1 in if ((case ((s168 s1710)) of Some (s1) => True | _ => False )) then (case s168 s1710 of (Some (s1)) => Some (( 0b000 :: 3 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s172 s1710)) of Some (s1) => True | _ => False )) then (case s172 s1710 of (Some (s1)) => Some (( 0b001 :: 3 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s176 s1710)) of Some (s1) => True | _ => False )) then (case s176 s1710 of (Some (s1)) => Some (( 0b010 :: 3 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s180 s1710)) of Some (s1) => True | _ => False )) then (case s180 s1710 of (Some (s1)) => Some (( 0b011 :: 3 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s184 s1710)) of Some (s1) => True | _ => False )) then (case s184 s1710 of (Some (s1)) => Some (( 0b100 :: 3 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s188 s1710)) of Some (s1) => True | _ => False )) then (case s188 s1710 of (Some (s1)) => Some (( 0b101 :: 3 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s192 s1710)) of Some (s1) => True | _ => False )) then (case s192 s1710 of (Some (s1)) => Some (( 0b110 :: 3 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s196 s1710)) of Some (s1) => True | _ => False )) then (case s196 s1710 of (Some (s1)) => Some (( 0b111 :: 3 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val init_base_regs : unit -> M unit\\ definition init_base_regs :: \ unit \((register_value),(unit),(exception))monad \ where \ init_base_regs _ = ( (((((((((((((((((((((((((((((write_reg x1_ref zero_reg \ write_reg x2_ref zero_reg) \ write_reg x3_ref zero_reg) \ write_reg x4_ref zero_reg) \ write_reg x5_ref zero_reg) \ write_reg x6_ref zero_reg) \ write_reg x7_ref zero_reg) \ write_reg x8_ref zero_reg) \ write_reg x9_ref zero_reg) \ write_reg x10_ref zero_reg) \ write_reg x11_ref zero_reg) \ write_reg x12_ref zero_reg) \ write_reg x13_ref zero_reg) \ write_reg x14_ref zero_reg) \ write_reg x15_ref zero_reg) \ write_reg x16_ref zero_reg) \ write_reg x17_ref zero_reg) \ write_reg x18_ref zero_reg) \ write_reg x19_ref zero_reg) \ write_reg x20_ref zero_reg) \ write_reg x21_ref zero_reg) \ write_reg x22_ref zero_reg) \ write_reg x23_ref zero_reg) \ write_reg x24_ref zero_reg) \ write_reg x25_ref zero_reg) \ write_reg x26_ref zero_reg) \ write_reg x27_ref zero_reg) \ write_reg x28_ref zero_reg) \ write_reg x29_ref zero_reg) \ write_reg x30_ref zero_reg) \ write_reg x31_ref zero_reg )\ \ \\ Retrieves the architectural PC value. This is not necessarily the value found in the PC register as extensions may choose to override this function. The value in the PC register is the absolute virtual address of the instruction to fetch. \\ \ \\val get_arch_pc : unit -> M (mword ty32)\\ definition get_arch_pc :: \ unit \((register_value),((32)Word.word),(exception))monad \ where \ get_arch_pc _ = ( (read_reg PC_ref :: ( 32 Word.word) M))\ \ \\val get_next_pc : unit -> M (mword ty32)\\ definition get_next_pc :: \ unit \((register_value),((32)Word.word),(exception))monad \ where \ get_next_pc _ = ( (read_reg nextPC_ref :: ( 32 Word.word) M))\ \ \\val set_next_pc : mword ty32 -> M unit\\ definition set_next_pc :: \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_next_pc pc = ( write_reg nextPC_ref pc )\ for pc :: "(32)Word.word " \ \\val tick_pc : unit -> M unit\\ definition tick_pc :: \ unit \((register_value),(unit),(exception))monad \ where \ tick_pc _ = ( (read_reg nextPC_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . write_reg PC_ref w__0)))\ \ \\val undefined_Misa : unit -> M Misa\\ definition undefined_Misa :: \ unit \((register_value),(Misa),(exception))monad \ where \ undefined_Misa _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Misa_bits = w__0 |)))))\ \ \\val Mk_Misa : mword ty32 -> Misa\\ definition Mk_Misa :: \(32)Word.word \ Misa \ where \ Mk_Misa v = ( (| Misa_bits = v |) )\ for v :: "(32)Word.word " \ \\val _get_Misa_bits : Misa -> mword ty32\\ definition get_Misa_bits :: \ Misa \(32)Word.word \ where \ get_Misa_bits v = ( (subrange_vec_dec(Misa_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Misa " \ \\val _set_Misa_bits : register_ref regstate register_value Misa -> mword ty32 -> M unit\\ definition set_Misa_bits :: \((regstate),(register_value),(Misa))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(32)Word.word " \ \\val _update_Misa_bits : Misa -> mword ty32 -> Misa\\ definition update_Misa_bits :: \ Misa \(32)Word.word \ Misa \ where \ update_Misa_bits v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(32)Word.word " \ \\val _update_Counteren_bits : Counteren -> mword ty32 -> Counteren\\ \ \\val _update_Counterin_bits : Counterin -> mword ty32 -> Counterin\\ \ \\val _update_Fcsr_bits : Fcsr -> mword ty32 -> Fcsr\\ \ \\val _update_Mcause_bits : Mcause -> mword ty32 -> Mcause\\ \ \\val _update_Medeleg_bits : Medeleg -> mword ty32 -> Medeleg\\ \ \\val _update_Minterrupts_bits : Minterrupts -> mword ty32 -> Minterrupts\\ \ \\val _update_Mstatus_bits : Mstatus -> mword ty32 -> Mstatus\\ \ \\val _update_Mstatush_bits : Mstatush -> mword ty32 -> Mstatush\\ \ \\val _update_Mtvec_bits : Mtvec -> mword ty32 -> Mtvec\\ \ \\val _update_PTE_Bits_bits : PTE_Bits -> mword ty8 -> PTE_Bits\\ \ \\val _update_Pmpcfg_ent_bits : Pmpcfg_ent -> mword ty8 -> Pmpcfg_ent\\ \ \\val _update_SV32_PTE_bits : SV32_PTE -> mword ty32 -> SV32_PTE\\ \ \\val _update_SV32_Paddr_bits : SV32_Paddr -> mword ty34 -> SV32_Paddr\\ \ \\val _update_SV32_Vaddr_bits : SV32_Vaddr -> mword ty32 -> SV32_Vaddr\\ \ \\val _update_SV39_PTE_bits : SV39_PTE -> mword ty64 -> SV39_PTE\\ \ \\val _update_SV39_Paddr_bits : SV39_Paddr -> mword ty56 -> SV39_Paddr\\ \ \\val _update_SV39_Vaddr_bits : SV39_Vaddr -> mword ty39 -> SV39_Vaddr\\ \ \\val _update_SV48_PTE_bits : SV48_PTE -> mword ty64 -> SV48_PTE\\ \ \\val _update_SV48_Paddr_bits : SV48_Paddr -> mword ty56 -> SV48_Paddr\\ \ \\val _update_SV48_Vaddr_bits : SV48_Vaddr -> mword ty48 -> SV48_Vaddr\\ \ \\val _update_Satp32_bits : Satp32 -> mword ty32 -> Satp32\\ \ \\val _update_Satp64_bits : Satp64 -> mword ty64 -> Satp64\\ \ \\val _update_Sedeleg_bits : Sedeleg -> mword ty32 -> Sedeleg\\ \ \\val _update_Sinterrupts_bits : Sinterrupts -> mword ty32 -> Sinterrupts\\ \ \\val _update_Sstatus_bits : Sstatus -> mword ty32 -> Sstatus\\ \ \\val _update_Uinterrupts_bits : Uinterrupts -> mword ty32 -> Uinterrupts\\ \ \\val _update_Ustatus_bits : Ustatus -> mword ty32 -> Ustatus\\ \ \\val _update_htif_cmd_bits : htif_cmd -> mword ty64 -> htif_cmd\\ \ \\val _get_Counteren_bits : Counteren -> mword ty32\\ \ \\val _get_Counterin_bits : Counterin -> mword ty32\\ \ \\val _get_Fcsr_bits : Fcsr -> mword ty32\\ \ \\val _get_Mcause_bits : Mcause -> mword ty32\\ \ \\val _get_Medeleg_bits : Medeleg -> mword ty32\\ \ \\val _get_Minterrupts_bits : Minterrupts -> mword ty32\\ \ \\val _get_Mstatus_bits : Mstatus -> mword ty32\\ \ \\val _get_Mstatush_bits : Mstatush -> mword ty32\\ \ \\val _get_Mtvec_bits : Mtvec -> mword ty32\\ \ \\val _get_PTE_Bits_bits : PTE_Bits -> mword ty8\\ \ \\val _get_Pmpcfg_ent_bits : Pmpcfg_ent -> mword ty8\\ \ \\val _get_SV32_PTE_bits : SV32_PTE -> mword ty32\\ \ \\val _get_SV32_Paddr_bits : SV32_Paddr -> mword ty34\\ \ \\val _get_SV32_Vaddr_bits : SV32_Vaddr -> mword ty32\\ \ \\val _get_SV39_PTE_bits : SV39_PTE -> mword ty64\\ \ \\val _get_SV39_Paddr_bits : SV39_Paddr -> mword ty56\\ \ \\val _get_SV39_Vaddr_bits : SV39_Vaddr -> mword ty39\\ \ \\val _get_SV48_PTE_bits : SV48_PTE -> mword ty64\\ \ \\val _get_SV48_Paddr_bits : SV48_Paddr -> mword ty56\\ \ \\val _get_SV48_Vaddr_bits : SV48_Vaddr -> mword ty48\\ \ \\val _get_Satp32_bits : Satp32 -> mword ty32\\ \ \\val _get_Satp64_bits : Satp64 -> mword ty64\\ \ \\val _get_Sedeleg_bits : Sedeleg -> mword ty32\\ \ \\val _get_Sinterrupts_bits : Sinterrupts -> mword ty32\\ \ \\val _get_Sstatus_bits : Sstatus -> mword ty32\\ \ \\val _get_Uinterrupts_bits : Uinterrupts -> mword ty32\\ \ \\val _get_Ustatus_bits : Ustatus -> mword ty32\\ \ \\val _get_htif_cmd_bits : htif_cmd -> mword ty64\\ \ \\val _set_Counteren_bits : register_ref regstate register_value Counteren -> mword ty32 -> M unit\\ \ \\val _set_Counterin_bits : register_ref regstate register_value Counterin -> mword ty32 -> M unit\\ \ \\val _set_Fcsr_bits : register_ref regstate register_value Fcsr -> mword ty32 -> M unit\\ \ \\val _set_Mcause_bits : register_ref regstate register_value Mcause -> mword ty32 -> M unit\\ \ \\val _set_Medeleg_bits : register_ref regstate register_value Medeleg -> mword ty32 -> M unit\\ \ \\val _set_Minterrupts_bits : register_ref regstate register_value Minterrupts -> mword ty32 -> M unit\\ \ \\val _set_Mstatus_bits : register_ref regstate register_value Mstatus -> mword ty32 -> M unit\\ \ \\val _set_Mstatush_bits : register_ref regstate register_value Mstatush -> mword ty32 -> M unit\\ \ \\val _set_Mtvec_bits : register_ref regstate register_value Mtvec -> mword ty32 -> M unit\\ \ \\val _set_PTE_Bits_bits : register_ref regstate register_value PTE_Bits -> mword ty8 -> M unit\\ \ \\val _set_Pmpcfg_ent_bits : register_ref regstate register_value Pmpcfg_ent -> mword ty8 -> M unit\\ \ \\val _set_SV32_PTE_bits : register_ref regstate register_value SV32_PTE -> mword ty32 -> M unit\\ \ \\val _set_SV32_Paddr_bits : register_ref regstate register_value SV32_Paddr -> mword ty34 -> M unit\\ \ \\val _set_SV32_Vaddr_bits : register_ref regstate register_value SV32_Vaddr -> mword ty32 -> M unit\\ \ \\val _set_SV39_PTE_bits : register_ref regstate register_value SV39_PTE -> mword ty64 -> M unit\\ \ \\val _set_SV39_Paddr_bits : register_ref regstate register_value SV39_Paddr -> mword ty56 -> M unit\\ \ \\val _set_SV39_Vaddr_bits : register_ref regstate register_value SV39_Vaddr -> mword ty39 -> M unit\\ \ \\val _set_SV48_PTE_bits : register_ref regstate register_value SV48_PTE -> mword ty64 -> M unit\\ \ \\val _set_SV48_Paddr_bits : register_ref regstate register_value SV48_Paddr -> mword ty56 -> M unit\\ \ \\val _set_SV48_Vaddr_bits : register_ref regstate register_value SV48_Vaddr -> mword ty48 -> M unit\\ \ \\val _set_Satp32_bits : register_ref regstate register_value Satp32 -> mword ty32 -> M unit\\ \ \\val _set_Satp64_bits : register_ref regstate register_value Satp64 -> mword ty64 -> M unit\\ \ \\val _set_Sedeleg_bits : register_ref regstate register_value Sedeleg -> mword ty32 -> M unit\\ \ \\val _set_Sinterrupts_bits : register_ref regstate register_value Sinterrupts -> mword ty32 -> M unit\\ \ \\val _set_Sstatus_bits : register_ref regstate register_value Sstatus -> mword ty32 -> M unit\\ \ \\val _set_Uinterrupts_bits : register_ref regstate register_value Uinterrupts -> mword ty32 -> M unit\\ \ \\val _set_Ustatus_bits : register_ref regstate register_value Ustatus -> mword ty32 -> M unit\\ \ \\val _set_htif_cmd_bits : register_ref regstate register_value htif_cmd -> mword ty64 -> M unit\\ \ \\val _get_Misa_A : Misa -> mword ty1\\ definition get_Misa_A :: \ Misa \(1)Word.word \ where \ get_Misa_A v = ( (subrange_vec_dec(Misa_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_A : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_A :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_A r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_A : Misa -> mword ty1 -> Misa\\ definition update_Misa_A :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_A v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_PTE_Bits_A : PTE_Bits -> mword ty1 -> PTE_Bits\\ \ \\val _update_Pmpcfg_ent_A : Pmpcfg_ent -> mword ty2 -> Pmpcfg_ent\\ \ \\val _get_PTE_Bits_A : PTE_Bits -> mword ty1\\ \ \\val _get_Pmpcfg_ent_A : Pmpcfg_ent -> mword ty2\\ \ \\val _set_PTE_Bits_A : register_ref regstate register_value PTE_Bits -> mword ty1 -> M unit\\ \ \\val _set_Pmpcfg_ent_A : register_ref regstate register_value Pmpcfg_ent -> mword ty2 -> M unit\\ \ \\val _get_Misa_B : Misa -> mword ty1\\ definition get_Misa_B :: \ Misa \(1)Word.word \ where \ get_Misa_B v = ( (subrange_vec_dec(Misa_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_B : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_B :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_B r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_B : Misa -> mword ty1 -> Misa\\ definition update_Misa_B :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_B v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_C : Misa -> mword ty1\\ definition get_Misa_C :: \ Misa \(1)Word.word \ where \ get_Misa_C v = ( (subrange_vec_dec(Misa_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_C : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_C :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_C r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 2 :: int)::ii) (( 2 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_C : Misa -> mword ty1 -> Misa\\ definition update_Misa_C :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_C v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_D : Misa -> mword ty1\\ definition get_Misa_D :: \ Misa \(1)Word.word \ where \ get_Misa_D v = ( (subrange_vec_dec(Misa_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_D : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_D :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_D r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 3 :: int)::ii) (( 3 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_D : Misa -> mword ty1 -> Misa\\ definition update_Misa_D :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_D v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_PTE_Bits_D : PTE_Bits -> mword ty1 -> PTE_Bits\\ \ \\val _get_PTE_Bits_D : PTE_Bits -> mword ty1\\ \ \\val _set_PTE_Bits_D : register_ref regstate register_value PTE_Bits -> mword ty1 -> M unit\\ \ \\val _get_Misa_E : Misa -> mword ty1\\ definition get_Misa_E :: \ Misa \(1)Word.word \ where \ get_Misa_E v = ( (subrange_vec_dec(Misa_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_E : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_E :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_E r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_E : Misa -> mword ty1 -> Misa\\ definition update_Misa_E :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_E v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_F : Misa -> mword ty1\\ definition get_Misa_F :: \ Misa \(1)Word.word \ where \ get_Misa_F v = ( (subrange_vec_dec(Misa_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_F : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_F :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_F r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_F : Misa -> mword ty1 -> Misa\\ definition update_Misa_F :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_F v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_G : Misa -> mword ty1\\ definition get_Misa_G :: \ Misa \(1)Word.word \ where \ get_Misa_G v = ( (subrange_vec_dec(Misa_bits v) (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_G : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_G :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_G r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 6 :: int)::ii) (( 6 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_G : Misa -> mword ty1 -> Misa\\ definition update_Misa_G :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_G v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 6 :: int)::ii) (( 6 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_PTE_Bits_G : PTE_Bits -> mword ty1 -> PTE_Bits\\ \ \\val _get_PTE_Bits_G : PTE_Bits -> mword ty1\\ \ \\val _set_PTE_Bits_G : register_ref regstate register_value PTE_Bits -> mword ty1 -> M unit\\ \ \\val _get_Misa_H : Misa -> mword ty1\\ definition get_Misa_H :: \ Misa \(1)Word.word \ where \ get_Misa_H v = ( (subrange_vec_dec(Misa_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_H : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_H :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_H r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 7 :: int)::ii) (( 7 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_H : Misa -> mword ty1 -> Misa\\ definition update_Misa_H :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_H v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_I : Misa -> mword ty1\\ definition get_Misa_I :: \ Misa \(1)Word.word \ where \ get_Misa_I v = ( (subrange_vec_dec(Misa_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_I : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_I :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_I r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 8 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_I : Misa -> mword ty1 -> Misa\\ definition update_Misa_I :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_I v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_J : Misa -> mword ty1\\ definition get_Misa_J :: \ Misa \(1)Word.word \ where \ get_Misa_J v = ( (subrange_vec_dec(Misa_bits v) (( 9 :: int)::ii) (( 9 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_J : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_J :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_J r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 9 :: int)::ii) (( 9 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_J : Misa -> mword ty1 -> Misa\\ definition update_Misa_J :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_J v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 9 :: int)::ii) (( 9 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_K : Misa -> mword ty1\\ definition get_Misa_K :: \ Misa \(1)Word.word \ where \ get_Misa_K v = ( (subrange_vec_dec(Misa_bits v) (( 10 :: int)::ii) (( 10 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_K : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_K :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_K r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 10 :: int)::ii) (( 10 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_K : Misa -> mword ty1 -> Misa\\ definition update_Misa_K :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_K v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 10 :: int)::ii) (( 10 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_L : Misa -> mword ty1\\ definition get_Misa_L :: \ Misa \(1)Word.word \ where \ get_Misa_L v = ( (subrange_vec_dec(Misa_bits v) (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_L : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_L :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_L r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 11 :: int)::ii) (( 11 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_L : Misa -> mword ty1 -> Misa\\ definition update_Misa_L :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_L v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 11 :: int)::ii) (( 11 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_Pmpcfg_ent_L : Pmpcfg_ent -> mword ty1 -> Pmpcfg_ent\\ \ \\val _get_Pmpcfg_ent_L : Pmpcfg_ent -> mword ty1\\ \ \\val _set_Pmpcfg_ent_L : register_ref regstate register_value Pmpcfg_ent -> mword ty1 -> M unit\\ \ \\val _get_Misa_M : Misa -> mword ty1\\ definition get_Misa_M :: \ Misa \(1)Word.word \ where \ get_Misa_M v = ( (subrange_vec_dec(Misa_bits v) (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_M : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_M :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_M r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 12 :: int)::ii) (( 12 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_M : Misa -> mword ty1 -> Misa\\ definition update_Misa_M :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_M v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 12 :: int)::ii) (( 12 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_MXL : Misa -> mword ty2\\ definition get_Misa_MXL :: \ Misa \(2)Word.word \ where \ get_Misa_MXL v = ( (subrange_vec_dec(Misa_bits v) (( 31 :: int)::ii) (( 30 :: int)::ii) :: 2 Word.word))\ for v :: " Misa " \ \\val _set_Misa_MXL : register_ref regstate register_value Misa -> mword ty2 -> M unit\\ definition set_Misa_MXL :: \((regstate),(register_value),(Misa))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_MXL r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 31 :: int)::ii) (( 30 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(2)Word.word " \ \\val _update_Misa_MXL : Misa -> mword ty2 -> Misa\\ definition update_Misa_MXL :: \ Misa \(2)Word.word \ Misa \ where \ update_Misa_MXL v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 31 :: int)::ii) (( 30 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(2)Word.word " \ \\val _get_Misa_N : Misa -> mword ty1\\ definition get_Misa_N :: \ Misa \(1)Word.word \ where \ get_Misa_N v = ( (subrange_vec_dec(Misa_bits v) (( 13 :: int)::ii) (( 13 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_N : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_N :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_N r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 13 :: int)::ii) (( 13 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_N : Misa -> mword ty1 -> Misa\\ definition update_Misa_N :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_N v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 13 :: int)::ii) (( 13 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_O : Misa -> mword ty1\\ definition get_Misa_O :: \ Misa \(1)Word.word \ where \ get_Misa_O v = ( (subrange_vec_dec(Misa_bits v) (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_O : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_O :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_O r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 14 :: int)::ii) (( 14 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_O : Misa -> mword ty1 -> Misa\\ definition update_Misa_O :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_O v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 14 :: int)::ii) (( 14 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_P : Misa -> mword ty1\\ definition get_Misa_P :: \ Misa \(1)Word.word \ where \ get_Misa_P v = ( (subrange_vec_dec(Misa_bits v) (( 15 :: int)::ii) (( 15 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_P : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_P :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_P r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 15 :: int)::ii) (( 15 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_P : Misa -> mword ty1 -> Misa\\ definition update_Misa_P :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_P v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 15 :: int)::ii) (( 15 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_Q : Misa -> mword ty1\\ definition get_Misa_Q :: \ Misa \(1)Word.word \ where \ get_Misa_Q v = ( (subrange_vec_dec(Misa_bits v) (( 16 :: int)::ii) (( 16 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_Q : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_Q :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_Q r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 16 :: int)::ii) (( 16 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_Q : Misa -> mword ty1 -> Misa\\ definition update_Misa_Q :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_Q v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 16 :: int)::ii) (( 16 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_R : Misa -> mword ty1\\ definition get_Misa_R :: \ Misa \(1)Word.word \ where \ get_Misa_R v = ( (subrange_vec_dec(Misa_bits v) (( 17 :: int)::ii) (( 17 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_R : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_R :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_R r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 17 :: int)::ii) (( 17 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_R : Misa -> mword ty1 -> Misa\\ definition update_Misa_R :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_R v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 17 :: int)::ii) (( 17 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_PTE_Bits_R : PTE_Bits -> mword ty1 -> PTE_Bits\\ \ \\val _update_Pmpcfg_ent_R : Pmpcfg_ent -> mword ty1 -> Pmpcfg_ent\\ \ \\val _get_PTE_Bits_R : PTE_Bits -> mword ty1\\ \ \\val _get_Pmpcfg_ent_R : Pmpcfg_ent -> mword ty1\\ \ \\val _set_PTE_Bits_R : register_ref regstate register_value PTE_Bits -> mword ty1 -> M unit\\ \ \\val _set_Pmpcfg_ent_R : register_ref regstate register_value Pmpcfg_ent -> mword ty1 -> M unit\\ \ \\val _get_Misa_S : Misa -> mword ty1\\ definition get_Misa_S :: \ Misa \(1)Word.word \ where \ get_Misa_S v = ( (subrange_vec_dec(Misa_bits v) (( 18 :: int)::ii) (( 18 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_S : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_S :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_S r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 18 :: int)::ii) (( 18 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_S : Misa -> mword ty1 -> Misa\\ definition update_Misa_S :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_S v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 18 :: int)::ii) (( 18 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_T : Misa -> mword ty1\\ definition get_Misa_T :: \ Misa \(1)Word.word \ where \ get_Misa_T v = ( (subrange_vec_dec(Misa_bits v) (( 19 :: int)::ii) (( 19 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_T : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_T :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_T r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 19 :: int)::ii) (( 19 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_T : Misa -> mword ty1 -> Misa\\ definition update_Misa_T :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_T v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 19 :: int)::ii) (( 19 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_U : Misa -> mword ty1\\ definition get_Misa_U :: \ Misa \(1)Word.word \ where \ get_Misa_U v = ( (subrange_vec_dec(Misa_bits v) (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_U : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_U :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_U r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 20 :: int)::ii) (( 20 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_U : Misa -> mword ty1 -> Misa\\ definition update_Misa_U :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_U v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 20 :: int)::ii) (( 20 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_PTE_Bits_U : PTE_Bits -> mword ty1 -> PTE_Bits\\ \ \\val _get_PTE_Bits_U : PTE_Bits -> mword ty1\\ \ \\val _set_PTE_Bits_U : register_ref regstate register_value PTE_Bits -> mword ty1 -> M unit\\ \ \\val _get_Misa_V : Misa -> mword ty1\\ definition get_Misa_V :: \ Misa \(1)Word.word \ where \ get_Misa_V v = ( (subrange_vec_dec(Misa_bits v) (( 21 :: int)::ii) (( 21 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_V : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_V :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_V r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 21 :: int)::ii) (( 21 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_V : Misa -> mword ty1 -> Misa\\ definition update_Misa_V :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_V v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 21 :: int)::ii) (( 21 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_PTE_Bits_V : PTE_Bits -> mword ty1 -> PTE_Bits\\ \ \\val _get_PTE_Bits_V : PTE_Bits -> mword ty1\\ \ \\val _set_PTE_Bits_V : register_ref regstate register_value PTE_Bits -> mword ty1 -> M unit\\ \ \\val _get_Misa_W : Misa -> mword ty1\\ definition get_Misa_W :: \ Misa \(1)Word.word \ where \ get_Misa_W v = ( (subrange_vec_dec(Misa_bits v) (( 22 :: int)::ii) (( 22 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_W : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_W :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_W r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 22 :: int)::ii) (( 22 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_W : Misa -> mword ty1 -> Misa\\ definition update_Misa_W :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_W v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 22 :: int)::ii) (( 22 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_PTE_Bits_W : PTE_Bits -> mword ty1 -> PTE_Bits\\ \ \\val _update_Pmpcfg_ent_W : Pmpcfg_ent -> mword ty1 -> Pmpcfg_ent\\ \ \\val _get_PTE_Bits_W : PTE_Bits -> mword ty1\\ \ \\val _get_Pmpcfg_ent_W : Pmpcfg_ent -> mword ty1\\ \ \\val _set_PTE_Bits_W : register_ref regstate register_value PTE_Bits -> mword ty1 -> M unit\\ \ \\val _set_Pmpcfg_ent_W : register_ref regstate register_value Pmpcfg_ent -> mword ty1 -> M unit\\ \ \\val _get_Misa_X : Misa -> mword ty1\\ definition get_Misa_X :: \ Misa \(1)Word.word \ where \ get_Misa_X v = ( (subrange_vec_dec(Misa_bits v) (( 23 :: int)::ii) (( 23 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_X : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_X :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_X r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 23 :: int)::ii) (( 23 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_X : Misa -> mword ty1 -> Misa\\ definition update_Misa_X :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_X v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 23 :: int)::ii) (( 23 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _update_PTE_Bits_X : PTE_Bits -> mword ty1 -> PTE_Bits\\ \ \\val _update_Pmpcfg_ent_X : Pmpcfg_ent -> mword ty1 -> Pmpcfg_ent\\ \ \\val _get_PTE_Bits_X : PTE_Bits -> mword ty1\\ \ \\val _get_Pmpcfg_ent_X : Pmpcfg_ent -> mword ty1\\ \ \\val _set_PTE_Bits_X : register_ref regstate register_value PTE_Bits -> mword ty1 -> M unit\\ \ \\val _set_Pmpcfg_ent_X : register_ref regstate register_value Pmpcfg_ent -> mword ty1 -> M unit\\ \ \\val _get_Misa_Y : Misa -> mword ty1\\ definition get_Misa_Y :: \ Misa \(1)Word.word \ where \ get_Misa_Y v = ( (subrange_vec_dec(Misa_bits v) (( 24 :: int)::ii) (( 24 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_Y : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_Y :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_Y r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 24 :: int)::ii) (( 24 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_Y : Misa -> mword ty1 -> Misa\\ definition update_Misa_Y :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_Y v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 24 :: int)::ii) (( 24 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val _get_Misa_Z : Misa -> mword ty1\\ definition get_Misa_Z :: \ Misa \(1)Word.word \ where \ get_Misa_Z v = ( (subrange_vec_dec(Misa_bits v) (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word))\ for v :: " Misa " \ \\val _set_Misa_Z : register_ref regstate register_value Misa -> mword ty1 -> M unit\\ definition set_Misa_Z :: \((regstate),(register_value),(Misa))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Misa_Z r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Misa_bits := ((update_subrange_vec_dec(Misa_bits r) (( 25 :: int)::ii) (( 25 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Misa))register_ref " and v :: "(1)Word.word " \ \\val _update_Misa_Z : Misa -> mword ty1 -> Misa\\ definition update_Misa_Z :: \ Misa \(1)Word.word \ Misa \ where \ update_Misa_Z v x = ( ( v (| Misa_bits := ((update_subrange_vec_dec(Misa_bits v) (( 25 :: int)::ii) (( 25 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Misa " and x :: "(1)Word.word " \ \\val ext_veto_disable_C : unit -> M bool\\ \ \\val legalize_misa : Misa -> mword ty32 -> M Misa\\ definition ext_veto_disable_C :: \ unit \((register_value),(bool),(exception))monad \ where \ ext_veto_disable_C _ = ( return False )\ definition legalize_misa :: \ Misa \(32)Word.word \((register_value),(Misa),(exception))monad \ where \ legalize_misa (m :: Misa) (v :: xlenbits) = ( if ((sys_enable_writable_misa () )) then (let v = (Mk_Misa v) in or_boolM (and_boolM (return (((((get_Misa_C v :: 1 Word.word)) = ( 0b0 :: 1 Word.word))))) (or_boolM ((read_reg nextPC_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return (((((access_vec_dec w__0 (( 1 :: int)::ii))) = B1)))))) ((ext_veto_disable_C () )))) (return ((\ ((sys_enable_rvc () ))))) \ ((\ (w__4 :: bool) . (let m = (if w__4 then m else update_Misa_C m ((get_Misa_C v :: 1 Word.word))) in return (if (((((\ ((sys_enable_fdext () )))) \ ((((((((get_Misa_D v :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ (((((get_Misa_F v :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))))))))) then m else update_Misa_D ((update_Misa_F m ((get_Misa_F v :: 1 Word.word)))) ((get_Misa_D v :: 1 Word.word))))))) else return m )\ for m :: " Misa " and v :: "(32)Word.word " \ \\val haveAtomics : unit -> M bool\\ definition haveAtomics :: \ unit \((register_value),(bool),(exception))monad \ where \ haveAtomics _ = ( read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_A w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))\ \ \\val haveRVC : unit -> M bool\\ definition haveRVC :: \ unit \((register_value),(bool),(exception))monad \ where \ haveRVC _ = ( read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_C w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))\ \ \\val haveMulDiv : unit -> M bool\\ definition haveMulDiv :: \ unit \((register_value),(bool),(exception))monad \ where \ haveMulDiv _ = ( read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_M w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))\ \ \\val haveSupMode : unit -> M bool\\ definition haveSupMode :: \ unit \((register_value),(bool),(exception))monad \ where \ haveSupMode _ = ( read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_S w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))\ \ \\val haveUsrMode : unit -> M bool\\ definition haveUsrMode :: \ unit \((register_value),(bool),(exception))monad \ where \ haveUsrMode _ = ( read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_U w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))\ \ \\val haveNExt : unit -> M bool\\ definition haveNExt :: \ unit \((register_value),(bool),(exception))monad \ where \ haveNExt _ = ( read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_N w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))\ \ \\val undefined_Mstatush : unit -> M Mstatush\\ definition undefined_Mstatush :: \ unit \((register_value),(Mstatush),(exception))monad \ where \ undefined_Mstatush _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Mstatush_bits = w__0 |)))))\ \ \\val Mk_Mstatush : mword ty32 -> Mstatush\\ definition Mk_Mstatush :: \(32)Word.word \ Mstatush \ where \ Mk_Mstatush v = ( (| Mstatush_bits = v |) )\ for v :: "(32)Word.word " definition get_Mstatush_bits :: \ Mstatush \(32)Word.word \ where \ get_Mstatush_bits v = ( (subrange_vec_dec(Mstatush_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Mstatush " definition set_Mstatush_bits :: \((regstate),(register_value),(Mstatush))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatush_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatush_bits := ((update_subrange_vec_dec(Mstatush_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatush))register_ref " and v :: "(32)Word.word " definition update_Mstatush_bits :: \ Mstatush \(32)Word.word \ Mstatush \ where \ update_Mstatush_bits v x = ( ( v (| Mstatush_bits := ((update_subrange_vec_dec(Mstatush_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatush " and x :: "(32)Word.word " \ \\val _get_Mstatush_MBE : Mstatush -> mword ty1\\ definition get_Mstatush_MBE :: \ Mstatush \(1)Word.word \ where \ get_Mstatush_MBE v = ( (subrange_vec_dec(Mstatush_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatush " \ \\val _set_Mstatush_MBE : register_ref regstate register_value Mstatush -> mword ty1 -> M unit\\ definition set_Mstatush_MBE :: \((regstate),(register_value),(Mstatush))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatush_MBE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatush_bits := ((update_subrange_vec_dec(Mstatush_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatush))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatush_MBE : Mstatush -> mword ty1 -> Mstatush\\ definition update_Mstatush_MBE :: \ Mstatush \(1)Word.word \ Mstatush \ where \ update_Mstatush_MBE v x = ( ( v (| Mstatush_bits := ((update_subrange_vec_dec(Mstatush_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatush " and x :: "(1)Word.word " \ \\val _get_Mstatush_SBE : Mstatush -> mword ty1\\ definition get_Mstatush_SBE :: \ Mstatush \(1)Word.word \ where \ get_Mstatush_SBE v = ( (subrange_vec_dec(Mstatush_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatush " \ \\val _set_Mstatush_SBE : register_ref regstate register_value Mstatush -> mword ty1 -> M unit\\ definition set_Mstatush_SBE :: \((regstate),(register_value),(Mstatush))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatush_SBE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatush_bits := ((update_subrange_vec_dec(Mstatush_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatush))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatush_SBE : Mstatush -> mword ty1 -> Mstatush\\ definition update_Mstatush_SBE :: \ Mstatush \(1)Word.word \ Mstatush \ where \ update_Mstatush_SBE v x = ( ( v (| Mstatush_bits := ((update_subrange_vec_dec(Mstatush_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatush " and x :: "(1)Word.word " \ \\val undefined_Mstatus : unit -> M Mstatus\\ definition undefined_Mstatus :: \ unit \((register_value),(Mstatus),(exception))monad \ where \ undefined_Mstatus _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Mstatus_bits = w__0 |)))))\ \ \\val Mk_Mstatus : mword ty32 -> Mstatus\\ definition Mk_Mstatus :: \(32)Word.word \ Mstatus \ where \ Mk_Mstatus v = ( (| Mstatus_bits = v |) )\ for v :: "(32)Word.word " definition get_Mstatus_bits :: \ Mstatus \(32)Word.word \ where \ get_Mstatus_bits v = ( (subrange_vec_dec(Mstatus_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Mstatus " definition set_Mstatus_bits :: \((regstate),(register_value),(Mstatus))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(32)Word.word " definition update_Mstatus_bits :: \ Mstatus \(32)Word.word \ Mstatus \ where \ update_Mstatus_bits v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(32)Word.word " \ \\val _get_Mstatus_FS : Mstatus -> mword ty2\\ definition get_Mstatus_FS :: \ Mstatus \(2)Word.word \ where \ get_Mstatus_FS v = ( (subrange_vec_dec(Mstatus_bits v) (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_FS : register_ref regstate register_value Mstatus -> mword ty2 -> M unit\\ definition set_Mstatus_FS :: \((regstate),(register_value),(Mstatus))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_FS r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 14 :: int)::ii) (( 13 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(2)Word.word " \ \\val _update_Mstatus_FS : Mstatus -> mword ty2 -> Mstatus\\ definition update_Mstatus_FS :: \ Mstatus \(2)Word.word \ Mstatus \ where \ update_Mstatus_FS v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 14 :: int)::ii) (( 13 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(2)Word.word " \ \\val _update_Sstatus_FS : Sstatus -> mword ty2 -> Sstatus\\ \ \\val _get_Sstatus_FS : Sstatus -> mword ty2\\ \ \\val _set_Sstatus_FS : register_ref regstate register_value Sstatus -> mword ty2 -> M unit\\ \ \\val _get_Mstatus_MIE : Mstatus -> mword ty1\\ definition get_Mstatus_MIE :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_MIE v = ( (subrange_vec_dec(Mstatus_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_MIE : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_MIE :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_MIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 3 :: int)::ii) (( 3 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_MIE : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_MIE :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_MIE v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _get_Mstatus_MPIE : Mstatus -> mword ty1\\ definition get_Mstatus_MPIE :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_MPIE v = ( (subrange_vec_dec(Mstatus_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_MPIE : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_MPIE :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_MPIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 7 :: int)::ii) (( 7 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_MPIE : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_MPIE :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_MPIE v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _get_Mstatus_MPP : Mstatus -> mword ty2\\ definition get_Mstatus_MPP :: \ Mstatus \(2)Word.word \ where \ get_Mstatus_MPP v = ( (subrange_vec_dec(Mstatus_bits v) (( 12 :: int)::ii) (( 11 :: int)::ii) :: 2 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_MPP : register_ref regstate register_value Mstatus -> mword ty2 -> M unit\\ definition set_Mstatus_MPP :: \((regstate),(register_value),(Mstatus))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_MPP r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 12 :: int)::ii) (( 11 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(2)Word.word " \ \\val _update_Mstatus_MPP : Mstatus -> mword ty2 -> Mstatus\\ definition update_Mstatus_MPP :: \ Mstatus \(2)Word.word \ Mstatus \ where \ update_Mstatus_MPP v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 12 :: int)::ii) (( 11 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(2)Word.word " \ \\val _get_Mstatus_MPRV : Mstatus -> mword ty1\\ definition get_Mstatus_MPRV :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_MPRV v = ( (subrange_vec_dec(Mstatus_bits v) (( 17 :: int)::ii) (( 17 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_MPRV : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_MPRV :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_MPRV r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 17 :: int)::ii) (( 17 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_MPRV : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_MPRV :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_MPRV v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 17 :: int)::ii) (( 17 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _get_Mstatus_MXR : Mstatus -> mword ty1\\ definition get_Mstatus_MXR :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_MXR v = ( (subrange_vec_dec(Mstatus_bits v) (( 19 :: int)::ii) (( 19 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_MXR : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_MXR :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_MXR r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 19 :: int)::ii) (( 19 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_MXR : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_MXR :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_MXR v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 19 :: int)::ii) (( 19 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _update_Sstatus_MXR : Sstatus -> mword ty1 -> Sstatus\\ \ \\val _get_Sstatus_MXR : Sstatus -> mword ty1\\ \ \\val _set_Sstatus_MXR : register_ref regstate register_value Sstatus -> mword ty1 -> M unit\\ \ \\val _get_Mstatus_SD : Mstatus -> mword ty1\\ definition get_Mstatus_SD :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_SD v = ( (subrange_vec_dec(Mstatus_bits v) (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_SD : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_SD :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_SD r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 31 :: int)::ii) (( 31 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_SD : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_SD :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_SD v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 31 :: int)::ii) (( 31 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _update_Sstatus_SD : Sstatus -> mword ty1 -> Sstatus\\ \ \\val _get_Sstatus_SD : Sstatus -> mword ty1\\ \ \\val _set_Sstatus_SD : register_ref regstate register_value Sstatus -> mword ty1 -> M unit\\ \ \\val _get_Mstatus_SIE : Mstatus -> mword ty1\\ definition get_Mstatus_SIE :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_SIE v = ( (subrange_vec_dec(Mstatus_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_SIE : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_SIE :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_SIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_SIE : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_SIE :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_SIE v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _update_Sstatus_SIE : Sstatus -> mword ty1 -> Sstatus\\ \ \\val _get_Sstatus_SIE : Sstatus -> mword ty1\\ \ \\val _set_Sstatus_SIE : register_ref regstate register_value Sstatus -> mword ty1 -> M unit\\ \ \\val _get_Mstatus_SPIE : Mstatus -> mword ty1\\ definition get_Mstatus_SPIE :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_SPIE v = ( (subrange_vec_dec(Mstatus_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_SPIE : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_SPIE :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_SPIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_SPIE : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_SPIE :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_SPIE v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _update_Sstatus_SPIE : Sstatus -> mword ty1 -> Sstatus\\ \ \\val _get_Sstatus_SPIE : Sstatus -> mword ty1\\ \ \\val _set_Sstatus_SPIE : register_ref regstate register_value Sstatus -> mword ty1 -> M unit\\ \ \\val _get_Mstatus_SPP : Mstatus -> mword ty1\\ definition get_Mstatus_SPP :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_SPP v = ( (subrange_vec_dec(Mstatus_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_SPP : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_SPP :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_SPP r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 8 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_SPP : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_SPP :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_SPP v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _update_Sstatus_SPP : Sstatus -> mword ty1 -> Sstatus\\ \ \\val _get_Sstatus_SPP : Sstatus -> mword ty1\\ \ \\val _set_Sstatus_SPP : register_ref regstate register_value Sstatus -> mword ty1 -> M unit\\ \ \\val _get_Mstatus_SUM : Mstatus -> mword ty1\\ definition get_Mstatus_SUM :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_SUM v = ( (subrange_vec_dec(Mstatus_bits v) (( 18 :: int)::ii) (( 18 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_SUM : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_SUM :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_SUM r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 18 :: int)::ii) (( 18 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_SUM : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_SUM :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_SUM v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 18 :: int)::ii) (( 18 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _update_Sstatus_SUM : Sstatus -> mword ty1 -> Sstatus\\ \ \\val _get_Sstatus_SUM : Sstatus -> mword ty1\\ \ \\val _set_Sstatus_SUM : register_ref regstate register_value Sstatus -> mword ty1 -> M unit\\ \ \\val _get_Mstatus_TSR : Mstatus -> mword ty1\\ definition get_Mstatus_TSR :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_TSR v = ( (subrange_vec_dec(Mstatus_bits v) (( 22 :: int)::ii) (( 22 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_TSR : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_TSR :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_TSR r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 22 :: int)::ii) (( 22 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_TSR : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_TSR :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_TSR v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 22 :: int)::ii) (( 22 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _get_Mstatus_TVM : Mstatus -> mword ty1\\ definition get_Mstatus_TVM :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_TVM v = ( (subrange_vec_dec(Mstatus_bits v) (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_TVM : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_TVM :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_TVM r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 20 :: int)::ii) (( 20 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_TVM : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_TVM :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_TVM v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 20 :: int)::ii) (( 20 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _get_Mstatus_TW : Mstatus -> mword ty1\\ definition get_Mstatus_TW :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_TW v = ( (subrange_vec_dec(Mstatus_bits v) (( 21 :: int)::ii) (( 21 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_TW : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_TW :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_TW r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 21 :: int)::ii) (( 21 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_TW : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_TW :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_TW v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 21 :: int)::ii) (( 21 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _get_Mstatus_UIE : Mstatus -> mword ty1\\ definition get_Mstatus_UIE :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_UIE v = ( (subrange_vec_dec(Mstatus_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_UIE : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_UIE :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_UIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_UIE : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_UIE :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_UIE v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _update_Sstatus_UIE : Sstatus -> mword ty1 -> Sstatus\\ \ \\val _update_Ustatus_UIE : Ustatus -> mword ty1 -> Ustatus\\ \ \\val _get_Sstatus_UIE : Sstatus -> mword ty1\\ \ \\val _get_Ustatus_UIE : Ustatus -> mword ty1\\ \ \\val _set_Sstatus_UIE : register_ref regstate register_value Sstatus -> mword ty1 -> M unit\\ \ \\val _set_Ustatus_UIE : register_ref regstate register_value Ustatus -> mword ty1 -> M unit\\ \ \\val _get_Mstatus_UPIE : Mstatus -> mword ty1\\ definition get_Mstatus_UPIE :: \ Mstatus \(1)Word.word \ where \ get_Mstatus_UPIE v = ( (subrange_vec_dec(Mstatus_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_UPIE : register_ref regstate register_value Mstatus -> mword ty1 -> M unit\\ definition set_Mstatus_UPIE :: \((regstate),(register_value),(Mstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_UPIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(1)Word.word " \ \\val _update_Mstatus_UPIE : Mstatus -> mword ty1 -> Mstatus\\ definition update_Mstatus_UPIE :: \ Mstatus \(1)Word.word \ Mstatus \ where \ update_Mstatus_UPIE v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(1)Word.word " \ \\val _update_Sstatus_UPIE : Sstatus -> mword ty1 -> Sstatus\\ \ \\val _update_Ustatus_UPIE : Ustatus -> mword ty1 -> Ustatus\\ \ \\val _get_Sstatus_UPIE : Sstatus -> mword ty1\\ \ \\val _get_Ustatus_UPIE : Ustatus -> mword ty1\\ \ \\val _set_Sstatus_UPIE : register_ref regstate register_value Sstatus -> mword ty1 -> M unit\\ \ \\val _set_Ustatus_UPIE : register_ref regstate register_value Ustatus -> mword ty1 -> M unit\\ \ \\val _get_Mstatus_XS : Mstatus -> mword ty2\\ definition get_Mstatus_XS :: \ Mstatus \(2)Word.word \ where \ get_Mstatus_XS v = ( (subrange_vec_dec(Mstatus_bits v) (( 16 :: int)::ii) (( 15 :: int)::ii) :: 2 Word.word))\ for v :: " Mstatus " \ \\val _set_Mstatus_XS : register_ref regstate register_value Mstatus -> mword ty2 -> M unit\\ definition set_Mstatus_XS :: \((regstate),(register_value),(Mstatus))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mstatus_XS r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits r) (( 16 :: int)::ii) (( 15 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mstatus))register_ref " and v :: "(2)Word.word " \ \\val _update_Mstatus_XS : Mstatus -> mword ty2 -> Mstatus\\ definition update_Mstatus_XS :: \ Mstatus \(2)Word.word \ Mstatus \ where \ update_Mstatus_XS v x = ( ( v (| Mstatus_bits := ((update_subrange_vec_dec(Mstatus_bits v) (( 16 :: int)::ii) (( 15 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mstatus " and x :: "(2)Word.word " \ \\val _update_Sstatus_XS : Sstatus -> mword ty2 -> Sstatus\\ \ \\val _get_Sstatus_XS : Sstatus -> mword ty2\\ \ \\val _set_Sstatus_XS : register_ref regstate register_value Sstatus -> mword ty2 -> M unit\\ \ \\val effectivePrivilege : AccessType unit -> Mstatus -> Privilege -> M Privilege\\ definition effectivePrivilege :: \(ext_access_type)AccessType \ Mstatus \ Privilege \((register_value),(Privilege),(exception))monad \ where \ effectivePrivilege (t :: ext_access_type AccessType) (m :: Mstatus) (priv :: Privilege) = ( if ((((((t \ (Execute () )))) \ (((((get_Mstatus_MPRV m :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))) then privLevel_of_bits ((get_Mstatus_MPP m :: 2 Word.word)) else return priv )\ for t :: "(ext_access_type)AccessType " and m :: " Mstatus " and priv :: " Privilege " \ \\val get_mstatus_SXL : Mstatus -> mword ty2\\ definition get_mstatus_SXL :: \ Mstatus \(2)Word.word \ where \ get_mstatus_SXL m = ( (arch_to_bits RV32 :: 2 Word.word))\ for m :: " Mstatus " \ \\val set_mstatus_SXL : Mstatus -> mword ty2 -> Mstatus\\ definition set_mstatus_SXL :: \ Mstatus \(2)Word.word \ Mstatus \ where \ set_mstatus_SXL (m :: Mstatus) (a :: arch_xlen) = ( m )\ for m :: " Mstatus " and a :: "(2)Word.word " \ \\val get_mstatus_UXL : Mstatus -> mword ty2\\ definition get_mstatus_UXL :: \ Mstatus \(2)Word.word \ where \ get_mstatus_UXL m = ( (arch_to_bits RV32 :: 2 Word.word))\ for m :: " Mstatus " \ \\val set_mstatus_UXL : Mstatus -> mword ty2 -> Mstatus\\ definition set_mstatus_UXL :: \ Mstatus \(2)Word.word \ Mstatus \ where \ set_mstatus_UXL (m :: Mstatus) (a :: arch_xlen) = ( m )\ for m :: " Mstatus " and a :: "(2)Word.word " \ \\val legalize_mstatus : Mstatus -> mword ty32 -> M Mstatus\\ definition legalize_mstatus :: \ Mstatus \(32)Word.word \((register_value),(Mstatus),(exception))monad \ where \ legalize_mstatus (o1 :: Mstatus) (v :: xlenbits) = ( (let (m :: Mstatus) = (Mk_Mstatus ((EXTZ (( 32 :: int)::ii) ((concat_vec ((subrange_vec_dec v (( 22 :: int)::ii) (( 11 :: int)::ii) :: 12 Word.word)) ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec ((subrange_vec_dec v (( 8 :: int)::ii) (( 7 :: int)::ii) :: 2 Word.word)) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ((subrange_vec_dec v (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) ((concat_vec ( 0b0 :: 1 Word.word) ((subrange_vec_dec v (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) :: 3 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 9 Word.word)) :: 11 Word.word)) :: 23 Word.word)) :: 32 Word.word))) in (let m = (update_Mstatus_XS m ((extStatus_to_bits Off :: 2 Word.word))) in or_boolM (extStatus_of_bits ((get_Mstatus_FS m :: 2 Word.word)) \ ((\ (w__0 :: ExtStatus) . return (((w__0 = Dirty)))))) (extStatus_of_bits ((get_Mstatus_XS m :: 2 Word.word)) \ ((\ (w__1 :: ExtStatus) . return (((w__1 = Dirty)))))) \ ((\ dirty . (let m = (update_Mstatus_SD m ((bool_to_bits dirty :: 1 Word.word))) in (let m = (set_mstatus_SXL m ((get_mstatus_SXL o1 :: 2 Word.word))) in (let m = (set_mstatus_UXL m ((get_mstatus_UXL o1 :: 2 Word.word))) in (let m = m in haveNExt () \ ((\ (w__2 :: bool) . (let m = (if ((\ w__2)) then (let m = (update_Mstatus_UPIE m ( 0b0 :: 1 Word.word)) in update_Mstatus_UIE m ( 0b0 :: 1 Word.word)) else m) in haveUsrMode () \ ((\ (w__3 :: bool) . return (if ((\ w__3)) then update_Mstatus_MPRV m ( 0b0 :: 1 Word.word) else m)))))))))))))))\ for o1 :: " Mstatus " and v :: "(32)Word.word " \ \\val cur_Architecture : unit -> M Architecture\\ definition cur_Architecture :: \ unit \((register_value),(Architecture),(exception))monad \ where \ cur_Architecture _ = ( read_reg cur_privilege_ref \ ((\ (w__0 :: Privilege) . (case w__0 of Machine => read_reg misa_ref \ ((\ (w__1 :: Misa) . return ((get_Misa_MXL w__1 :: 2 Word.word)))) | Supervisor => read_reg mstatus_ref \ ((\ (w__2 :: Mstatus) . return ((get_mstatus_SXL w__2 :: 2 Word.word)))) | User => read_reg mstatus_ref \ ((\ (w__3 :: Mstatus) . return ((get_mstatus_UXL w__3 :: 2 Word.word)))) ) \ ((\ (a :: arch_xlen) . (case ((architecture a)) of Some (a) => return a | None => internal_error (''Invalid current architecture'') ))))))\ \ \\val in32BitMode : unit -> M bool\\ definition in32BitMode :: \ unit \((register_value),(bool),(exception))monad \ where \ in32BitMode _ = ( cur_Architecture () \ ((\ (w__0 :: Architecture) . return (((w__0 = RV32))))))\ \ \\val haveFExt : unit -> M bool\\ definition haveFExt :: \ unit \((register_value),(bool),(exception))monad \ where \ haveFExt _ = ( and_boolM (read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_F w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))) (read_reg mstatus_ref \ ((\ (w__1 :: Mstatus) . return (((((get_Mstatus_FS w__1 :: 2 Word.word)) \ ( 0b00 :: 2 Word.word))))))))\ \ \\val haveDExt : unit -> M bool\\ definition haveDExt :: \ unit \((register_value),(bool),(exception))monad \ where \ haveDExt _ = ( and_boolM (read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_D w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))) (read_reg mstatus_ref \ ((\ (w__1 :: Mstatus) . return (((((get_Mstatus_FS w__1 :: 2 Word.word)) \ ( 0b00 :: 2 Word.word))))))))\ \ \\val undefined_Minterrupts : unit -> M Minterrupts\\ definition undefined_Minterrupts :: \ unit \((register_value),(Minterrupts),(exception))monad \ where \ undefined_Minterrupts _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Minterrupts_bits = w__0 |)))))\ \ \\val Mk_Minterrupts : mword ty32 -> Minterrupts\\ definition Mk_Minterrupts :: \(32)Word.word \ Minterrupts \ where \ Mk_Minterrupts v = ( (| Minterrupts_bits = v |) )\ for v :: "(32)Word.word " definition get_Minterrupts_bits :: \ Minterrupts \(32)Word.word \ where \ get_Minterrupts_bits v = ( (subrange_vec_dec(Minterrupts_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Minterrupts " definition set_Minterrupts_bits :: \((regstate),(register_value),(Minterrupts))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(32)Word.word " definition update_Minterrupts_bits :: \ Minterrupts \(32)Word.word \ Minterrupts \ where \ update_Minterrupts_bits v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(32)Word.word " \ \\val _get_Minterrupts_MEI : Minterrupts -> mword ty1\\ definition get_Minterrupts_MEI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_MEI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_MEI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_MEI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_MEI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 11 :: int)::ii) (( 11 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_MEI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_MEI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_MEI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 11 :: int)::ii) (( 11 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _get_Minterrupts_MSI : Minterrupts -> mword ty1\\ definition get_Minterrupts_MSI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_MSI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_MSI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_MSI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_MSI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 3 :: int)::ii) (( 3 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_MSI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_MSI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_MSI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _get_Minterrupts_MTI : Minterrupts -> mword ty1\\ definition get_Minterrupts_MTI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_MTI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_MTI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_MTI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_MTI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 7 :: int)::ii) (( 7 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_MTI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_MTI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_MTI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _get_Minterrupts_SEI : Minterrupts -> mword ty1\\ definition get_Minterrupts_SEI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_SEI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 9 :: int)::ii) (( 9 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_SEI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_SEI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_SEI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 9 :: int)::ii) (( 9 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_SEI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_SEI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_SEI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 9 :: int)::ii) (( 9 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _update_Sinterrupts_SEI : Sinterrupts -> mword ty1 -> Sinterrupts\\ \ \\val _get_Sinterrupts_SEI : Sinterrupts -> mword ty1\\ \ \\val _set_Sinterrupts_SEI : register_ref regstate register_value Sinterrupts -> mword ty1 -> M unit\\ \ \\val _get_Minterrupts_SSI : Minterrupts -> mword ty1\\ definition get_Minterrupts_SSI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_SSI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_SSI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_SSI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_SSI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_SSI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_SSI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_SSI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _update_Sinterrupts_SSI : Sinterrupts -> mword ty1 -> Sinterrupts\\ \ \\val _get_Sinterrupts_SSI : Sinterrupts -> mword ty1\\ \ \\val _set_Sinterrupts_SSI : register_ref regstate register_value Sinterrupts -> mword ty1 -> M unit\\ \ \\val _get_Minterrupts_STI : Minterrupts -> mword ty1\\ definition get_Minterrupts_STI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_STI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_STI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_STI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_STI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_STI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_STI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_STI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _update_Sinterrupts_STI : Sinterrupts -> mword ty1 -> Sinterrupts\\ \ \\val _get_Sinterrupts_STI : Sinterrupts -> mword ty1\\ \ \\val _set_Sinterrupts_STI : register_ref regstate register_value Sinterrupts -> mword ty1 -> M unit\\ \ \\val _get_Minterrupts_UEI : Minterrupts -> mword ty1\\ definition get_Minterrupts_UEI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_UEI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_UEI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_UEI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_UEI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 8 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_UEI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_UEI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_UEI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _update_Sinterrupts_UEI : Sinterrupts -> mword ty1 -> Sinterrupts\\ \ \\val _update_Uinterrupts_UEI : Uinterrupts -> mword ty1 -> Uinterrupts\\ \ \\val _get_Sinterrupts_UEI : Sinterrupts -> mword ty1\\ \ \\val _get_Uinterrupts_UEI : Uinterrupts -> mword ty1\\ \ \\val _set_Sinterrupts_UEI : register_ref regstate register_value Sinterrupts -> mword ty1 -> M unit\\ \ \\val _set_Uinterrupts_UEI : register_ref regstate register_value Uinterrupts -> mword ty1 -> M unit\\ \ \\val _get_Minterrupts_USI : Minterrupts -> mword ty1\\ definition get_Minterrupts_USI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_USI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_USI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_USI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_USI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_USI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_USI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_USI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _update_Sinterrupts_USI : Sinterrupts -> mword ty1 -> Sinterrupts\\ \ \\val _update_Uinterrupts_USI : Uinterrupts -> mword ty1 -> Uinterrupts\\ \ \\val _get_Sinterrupts_USI : Sinterrupts -> mword ty1\\ \ \\val _get_Uinterrupts_USI : Uinterrupts -> mword ty1\\ \ \\val _set_Sinterrupts_USI : register_ref regstate register_value Sinterrupts -> mword ty1 -> M unit\\ \ \\val _set_Uinterrupts_USI : register_ref regstate register_value Uinterrupts -> mword ty1 -> M unit\\ \ \\val _get_Minterrupts_UTI : Minterrupts -> mword ty1\\ definition get_Minterrupts_UTI :: \ Minterrupts \(1)Word.word \ where \ get_Minterrupts_UTI v = ( (subrange_vec_dec(Minterrupts_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Minterrupts " \ \\val _set_Minterrupts_UTI : register_ref regstate register_value Minterrupts -> mword ty1 -> M unit\\ definition set_Minterrupts_UTI :: \((regstate),(register_value),(Minterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Minterrupts_UTI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Minterrupts))register_ref " and v :: "(1)Word.word " \ \\val _update_Minterrupts_UTI : Minterrupts -> mword ty1 -> Minterrupts\\ definition update_Minterrupts_UTI :: \ Minterrupts \(1)Word.word \ Minterrupts \ where \ update_Minterrupts_UTI v x = ( ( v (| Minterrupts_bits := ((update_subrange_vec_dec(Minterrupts_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Minterrupts " and x :: "(1)Word.word " \ \\val _update_Sinterrupts_UTI : Sinterrupts -> mword ty1 -> Sinterrupts\\ \ \\val _update_Uinterrupts_UTI : Uinterrupts -> mword ty1 -> Uinterrupts\\ \ \\val _get_Sinterrupts_UTI : Sinterrupts -> mword ty1\\ \ \\val _get_Uinterrupts_UTI : Uinterrupts -> mword ty1\\ \ \\val _set_Sinterrupts_UTI : register_ref regstate register_value Sinterrupts -> mword ty1 -> M unit\\ \ \\val _set_Uinterrupts_UTI : register_ref regstate register_value Uinterrupts -> mword ty1 -> M unit\\ \ \\val legalize_mip : Minterrupts -> mword ty32 -> M Minterrupts\\ definition legalize_mip :: \ Minterrupts \(32)Word.word \((register_value),(Minterrupts),(exception))monad \ where \ legalize_mip (o1 :: Minterrupts) (v :: xlenbits) = ( (let v = (Mk_Minterrupts v) in (let m = (update_Minterrupts_SEI o1 ((get_Minterrupts_SEI v :: 1 Word.word))) in (let m = (update_Minterrupts_STI m ((get_Minterrupts_STI v :: 1 Word.word))) in (let m = (update_Minterrupts_SSI m ((get_Minterrupts_SSI v :: 1 Word.word))) in and_boolM ((haveUsrMode () )) ((haveNExt () )) \ ((\ (w__2 :: bool) . return (if w__2 then (let m = (update_Minterrupts_UEI m ((get_Minterrupts_UEI v :: 1 Word.word))) in (let m = (update_Minterrupts_UTI m ((get_Minterrupts_UTI v :: 1 Word.word))) in update_Minterrupts_USI m ((get_Minterrupts_USI v :: 1 Word.word)))) else m))))))))\ for o1 :: " Minterrupts " and v :: "(32)Word.word " \ \\val legalize_mie : Minterrupts -> mword ty32 -> M Minterrupts\\ definition legalize_mie :: \ Minterrupts \(32)Word.word \((register_value),(Minterrupts),(exception))monad \ where \ legalize_mie (o1 :: Minterrupts) (v :: xlenbits) = ( (let v = (Mk_Minterrupts v) in (let m = (update_Minterrupts_MEI o1 ((get_Minterrupts_MEI v :: 1 Word.word))) in (let m = (update_Minterrupts_MTI m ((get_Minterrupts_MTI v :: 1 Word.word))) in (let m = (update_Minterrupts_MSI m ((get_Minterrupts_MSI v :: 1 Word.word))) in (let m = (update_Minterrupts_SEI m ((get_Minterrupts_SEI v :: 1 Word.word))) in (let m = (update_Minterrupts_STI m ((get_Minterrupts_STI v :: 1 Word.word))) in (let m = (update_Minterrupts_SSI m ((get_Minterrupts_SSI v :: 1 Word.word))) in and_boolM ((haveUsrMode () )) ((haveNExt () )) \ ((\ (w__2 :: bool) . return (if w__2 then (let m = (update_Minterrupts_UEI m ((get_Minterrupts_UEI v :: 1 Word.word))) in (let m = (update_Minterrupts_UTI m ((get_Minterrupts_UTI v :: 1 Word.word))) in update_Minterrupts_USI m ((get_Minterrupts_USI v :: 1 Word.word)))) else m)))))))))))\ for o1 :: " Minterrupts " and v :: "(32)Word.word " \ \\val legalize_mideleg : Minterrupts -> mword ty32 -> Minterrupts\\ definition legalize_mideleg :: \ Minterrupts \(32)Word.word \ Minterrupts \ where \ legalize_mideleg (o1 :: Minterrupts) (v :: xlenbits) = ( (let m = (Mk_Minterrupts v) in (let m = (update_Minterrupts_MEI m ( 0b0 :: 1 Word.word)) in (let m = (update_Minterrupts_MTI m ( 0b0 :: 1 Word.word)) in update_Minterrupts_MSI m ( 0b0 :: 1 Word.word)))))\ for o1 :: " Minterrupts " and v :: "(32)Word.word " \ \\val undefined_Medeleg : unit -> M Medeleg\\ definition undefined_Medeleg :: \ unit \((register_value),(Medeleg),(exception))monad \ where \ undefined_Medeleg _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Medeleg_bits = w__0 |)))))\ \ \\val Mk_Medeleg : mword ty32 -> Medeleg\\ definition Mk_Medeleg :: \(32)Word.word \ Medeleg \ where \ Mk_Medeleg v = ( (| Medeleg_bits = v |) )\ for v :: "(32)Word.word " definition get_Medeleg_bits :: \ Medeleg \(32)Word.word \ where \ get_Medeleg_bits v = ( (subrange_vec_dec(Medeleg_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Medeleg " definition set_Medeleg_bits :: \((regstate),(register_value),(Medeleg))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(32)Word.word " definition update_Medeleg_bits :: \ Medeleg \(32)Word.word \ Medeleg \ where \ update_Medeleg_bits v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(32)Word.word " \ \\val _get_Medeleg_Breakpoint : Medeleg -> mword ty1\\ definition get_Medeleg_Breakpoint :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_Breakpoint v = ( (subrange_vec_dec(Medeleg_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_Breakpoint : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_Breakpoint :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_Breakpoint r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 3 :: int)::ii) (( 3 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_Breakpoint : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_Breakpoint :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_Breakpoint v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_Breakpoint : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_Breakpoint : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_Breakpoint : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val _get_Medeleg_Fetch_Access_Fault : Medeleg -> mword ty1\\ definition get_Medeleg_Fetch_Access_Fault :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_Fetch_Access_Fault v = ( (subrange_vec_dec(Medeleg_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_Fetch_Access_Fault : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_Fetch_Access_Fault :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_Fetch_Access_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_Fetch_Access_Fault : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_Fetch_Access_Fault :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_Fetch_Access_Fault v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_Fetch_Access_Fault : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_Fetch_Access_Fault : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_Fetch_Access_Fault : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val _get_Medeleg_Fetch_Addr_Align : Medeleg -> mword ty1\\ definition get_Medeleg_Fetch_Addr_Align :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_Fetch_Addr_Align v = ( (subrange_vec_dec(Medeleg_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_Fetch_Addr_Align : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_Fetch_Addr_Align :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_Fetch_Addr_Align r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_Fetch_Addr_Align : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_Fetch_Addr_Align :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_Fetch_Addr_Align v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_Fetch_Addr_Align : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_Fetch_Addr_Align : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_Fetch_Addr_Align : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val _get_Medeleg_Fetch_Page_Fault : Medeleg -> mword ty1\\ definition get_Medeleg_Fetch_Page_Fault :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_Fetch_Page_Fault v = ( (subrange_vec_dec(Medeleg_bits v) (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_Fetch_Page_Fault : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_Fetch_Page_Fault :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_Fetch_Page_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 12 :: int)::ii) (( 12 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_Fetch_Page_Fault : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_Fetch_Page_Fault :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_Fetch_Page_Fault v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 12 :: int)::ii) (( 12 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _get_Medeleg_Illegal_Instr : Medeleg -> mword ty1\\ definition get_Medeleg_Illegal_Instr :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_Illegal_Instr v = ( (subrange_vec_dec(Medeleg_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_Illegal_Instr : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_Illegal_Instr :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_Illegal_Instr r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 2 :: int)::ii) (( 2 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_Illegal_Instr : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_Illegal_Instr :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_Illegal_Instr v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_Illegal_Instr : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_Illegal_Instr : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_Illegal_Instr : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val _get_Medeleg_Load_Access_Fault : Medeleg -> mword ty1\\ definition get_Medeleg_Load_Access_Fault :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_Load_Access_Fault v = ( (subrange_vec_dec(Medeleg_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_Load_Access_Fault : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_Load_Access_Fault :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_Load_Access_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_Load_Access_Fault : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_Load_Access_Fault :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_Load_Access_Fault v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_Load_Access_Fault : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_Load_Access_Fault : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_Load_Access_Fault : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val _get_Medeleg_Load_Addr_Align : Medeleg -> mword ty1\\ definition get_Medeleg_Load_Addr_Align :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_Load_Addr_Align v = ( (subrange_vec_dec(Medeleg_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_Load_Addr_Align : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_Load_Addr_Align :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_Load_Addr_Align r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_Load_Addr_Align : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_Load_Addr_Align :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_Load_Addr_Align v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_Load_Addr_Align : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_Load_Addr_Align : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_Load_Addr_Align : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val _get_Medeleg_Load_Page_Fault : Medeleg -> mword ty1\\ definition get_Medeleg_Load_Page_Fault :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_Load_Page_Fault v = ( (subrange_vec_dec(Medeleg_bits v) (( 13 :: int)::ii) (( 13 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_Load_Page_Fault : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_Load_Page_Fault :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_Load_Page_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 13 :: int)::ii) (( 13 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_Load_Page_Fault : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_Load_Page_Fault :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_Load_Page_Fault v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 13 :: int)::ii) (( 13 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _get_Medeleg_MEnvCall : Medeleg -> mword ty1\\ definition get_Medeleg_MEnvCall :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_MEnvCall v = ( (subrange_vec_dec(Medeleg_bits v) (( 10 :: int)::ii) (( 10 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_MEnvCall : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_MEnvCall :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_MEnvCall r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 10 :: int)::ii) (( 10 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_MEnvCall : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_MEnvCall :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_MEnvCall v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 10 :: int)::ii) (( 10 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _get_Medeleg_SAMO_Access_Fault : Medeleg -> mword ty1\\ definition get_Medeleg_SAMO_Access_Fault :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_SAMO_Access_Fault v = ( (subrange_vec_dec(Medeleg_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_SAMO_Access_Fault : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_SAMO_Access_Fault :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_SAMO_Access_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 7 :: int)::ii) (( 7 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_SAMO_Access_Fault : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_SAMO_Access_Fault :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_SAMO_Access_Fault v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_SAMO_Access_Fault : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_SAMO_Access_Fault : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_SAMO_Access_Fault : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val _get_Medeleg_SAMO_Addr_Align : Medeleg -> mword ty1\\ definition get_Medeleg_SAMO_Addr_Align :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_SAMO_Addr_Align v = ( (subrange_vec_dec(Medeleg_bits v) (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_SAMO_Addr_Align : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_SAMO_Addr_Align :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_SAMO_Addr_Align r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 6 :: int)::ii) (( 6 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_SAMO_Addr_Align : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_SAMO_Addr_Align :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_SAMO_Addr_Align v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 6 :: int)::ii) (( 6 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_SAMO_Addr_Align : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_SAMO_Addr_Align : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_SAMO_Addr_Align : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val _get_Medeleg_SAMO_Page_Fault : Medeleg -> mword ty1\\ definition get_Medeleg_SAMO_Page_Fault :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_SAMO_Page_Fault v = ( (subrange_vec_dec(Medeleg_bits v) (( 15 :: int)::ii) (( 15 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_SAMO_Page_Fault : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_SAMO_Page_Fault :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_SAMO_Page_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 15 :: int)::ii) (( 15 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_SAMO_Page_Fault : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_SAMO_Page_Fault :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_SAMO_Page_Fault v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 15 :: int)::ii) (( 15 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _get_Medeleg_SEnvCall : Medeleg -> mword ty1\\ definition get_Medeleg_SEnvCall :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_SEnvCall v = ( (subrange_vec_dec(Medeleg_bits v) (( 9 :: int)::ii) (( 9 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_SEnvCall : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_SEnvCall :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_SEnvCall r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 9 :: int)::ii) (( 9 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_SEnvCall : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_SEnvCall :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_SEnvCall v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 9 :: int)::ii) (( 9 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _get_Medeleg_UEnvCall : Medeleg -> mword ty1\\ definition get_Medeleg_UEnvCall :: \ Medeleg \(1)Word.word \ where \ get_Medeleg_UEnvCall v = ( (subrange_vec_dec(Medeleg_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word))\ for v :: " Medeleg " \ \\val _set_Medeleg_UEnvCall : register_ref regstate register_value Medeleg -> mword ty1 -> M unit\\ definition set_Medeleg_UEnvCall :: \((regstate),(register_value),(Medeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Medeleg_UEnvCall r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits r) (( 8 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Medeleg))register_ref " and v :: "(1)Word.word " \ \\val _update_Medeleg_UEnvCall : Medeleg -> mword ty1 -> Medeleg\\ definition update_Medeleg_UEnvCall :: \ Medeleg \(1)Word.word \ Medeleg \ where \ update_Medeleg_UEnvCall v x = ( ( v (| Medeleg_bits := ((update_subrange_vec_dec(Medeleg_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Medeleg " and x :: "(1)Word.word " \ \\val _update_Sedeleg_UEnvCall : Sedeleg -> mword ty1 -> Sedeleg\\ \ \\val _get_Sedeleg_UEnvCall : Sedeleg -> mword ty1\\ \ \\val _set_Sedeleg_UEnvCall : register_ref regstate register_value Sedeleg -> mword ty1 -> M unit\\ \ \\val legalize_medeleg : Medeleg -> mword ty32 -> Medeleg\\ definition legalize_medeleg :: \ Medeleg \(32)Word.word \ Medeleg \ where \ legalize_medeleg (o1 :: Medeleg) (v :: xlenbits) = ( (let m = (Mk_Medeleg v) in update_Medeleg_MEnvCall m ( 0b0 :: 1 Word.word)))\ for o1 :: " Medeleg " and v :: "(32)Word.word " \ \\val undefined_Mtvec : unit -> M Mtvec\\ definition undefined_Mtvec :: \ unit \((register_value),(Mtvec),(exception))monad \ where \ undefined_Mtvec _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Mtvec_bits = w__0 |)))))\ \ \\val Mk_Mtvec : mword ty32 -> Mtvec\\ definition Mk_Mtvec :: \(32)Word.word \ Mtvec \ where \ Mk_Mtvec v = ( (| Mtvec_bits = v |) )\ for v :: "(32)Word.word " definition get_Mtvec_bits :: \ Mtvec \(32)Word.word \ where \ get_Mtvec_bits v = ( (subrange_vec_dec(Mtvec_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Mtvec " definition set_Mtvec_bits :: \((regstate),(register_value),(Mtvec))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mtvec_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mtvec_bits := ((update_subrange_vec_dec(Mtvec_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mtvec))register_ref " and v :: "(32)Word.word " definition update_Mtvec_bits :: \ Mtvec \(32)Word.word \ Mtvec \ where \ update_Mtvec_bits v x = ( ( v (| Mtvec_bits := ((update_subrange_vec_dec(Mtvec_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mtvec " and x :: "(32)Word.word " \ \\val _get_Mtvec_Base : Mtvec -> mword ty30\\ definition get_Mtvec_Base :: \ Mtvec \(30)Word.word \ where \ get_Mtvec_Base v = ( (subrange_vec_dec(Mtvec_bits v) (( 31 :: int)::ii) (( 2 :: int)::ii) :: 30 Word.word))\ for v :: " Mtvec " \ \\val _set_Mtvec_Base : register_ref regstate register_value Mtvec -> mword ty30 -> M unit\\ definition set_Mtvec_Base :: \((regstate),(register_value),(Mtvec))register_ref \(30)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mtvec_Base r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mtvec_bits := ((update_subrange_vec_dec(Mtvec_bits r) (( 31 :: int)::ii) (( 2 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mtvec))register_ref " and v :: "(30)Word.word " \ \\val _update_Mtvec_Base : Mtvec -> mword ty30 -> Mtvec\\ definition update_Mtvec_Base :: \ Mtvec \(30)Word.word \ Mtvec \ where \ update_Mtvec_Base v x = ( ( v (| Mtvec_bits := ((update_subrange_vec_dec(Mtvec_bits v) (( 31 :: int)::ii) (( 2 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mtvec " and x :: "(30)Word.word " \ \\val _get_Mtvec_Mode : Mtvec -> mword ty2\\ definition get_Mtvec_Mode :: \ Mtvec \(2)Word.word \ where \ get_Mtvec_Mode v = ( (subrange_vec_dec(Mtvec_bits v) (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word))\ for v :: " Mtvec " \ \\val _set_Mtvec_Mode : register_ref regstate register_value Mtvec -> mword ty2 -> M unit\\ definition set_Mtvec_Mode :: \((regstate),(register_value),(Mtvec))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mtvec_Mode r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mtvec_bits := ((update_subrange_vec_dec(Mtvec_bits r) (( 1 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mtvec))register_ref " and v :: "(2)Word.word " \ \\val _update_Mtvec_Mode : Mtvec -> mword ty2 -> Mtvec\\ definition update_Mtvec_Mode :: \ Mtvec \(2)Word.word \ Mtvec \ where \ update_Mtvec_Mode v x = ( ( v (| Mtvec_bits := ((update_subrange_vec_dec(Mtvec_bits v) (( 1 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mtvec " and x :: "(2)Word.word " \ \\val _update_Satp32_Mode : Satp32 -> mword ty1 -> Satp32\\ \ \\val _update_Satp64_Mode : Satp64 -> mword ty4 -> Satp64\\ \ \\val _get_Satp32_Mode : Satp32 -> mword ty1\\ \ \\val _get_Satp64_Mode : Satp64 -> mword ty4\\ \ \\val _set_Satp32_Mode : register_ref regstate register_value Satp32 -> mword ty1 -> M unit\\ \ \\val _set_Satp64_Mode : register_ref regstate register_value Satp64 -> mword ty4 -> M unit\\ \ \\val legalize_tvec : Mtvec -> mword ty32 -> Mtvec\\ definition legalize_tvec :: \ Mtvec \(32)Word.word \ Mtvec \ where \ legalize_tvec (o1 :: Mtvec) (v :: xlenbits) = ( (let v = (Mk_Mtvec v) in (case ((trapVectorMode_of_bits ((get_Mtvec_Mode v :: 2 Word.word)))) of TV_Direct => v | TV_Vector => v | _ => update_Mtvec_Mode v ((get_Mtvec_Mode o1 :: 2 Word.word)) )))\ for o1 :: " Mtvec " and v :: "(32)Word.word " \ \\val undefined_Mcause : unit -> M Mcause\\ definition undefined_Mcause :: \ unit \((register_value),(Mcause),(exception))monad \ where \ undefined_Mcause _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Mcause_bits = w__0 |)))))\ \ \\val Mk_Mcause : mword ty32 -> Mcause\\ definition Mk_Mcause :: \(32)Word.word \ Mcause \ where \ Mk_Mcause v = ( (| Mcause_bits = v |) )\ for v :: "(32)Word.word " definition get_Mcause_bits :: \ Mcause \(32)Word.word \ where \ get_Mcause_bits v = ( (subrange_vec_dec(Mcause_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Mcause " definition set_Mcause_bits :: \((regstate),(register_value),(Mcause))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mcause_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mcause_bits := ((update_subrange_vec_dec(Mcause_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mcause))register_ref " and v :: "(32)Word.word " definition update_Mcause_bits :: \ Mcause \(32)Word.word \ Mcause \ where \ update_Mcause_bits v x = ( ( v (| Mcause_bits := ((update_subrange_vec_dec(Mcause_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mcause " and x :: "(32)Word.word " \ \\val _get_Mcause_Cause : Mcause -> mword ty31\\ definition get_Mcause_Cause :: \ Mcause \(31)Word.word \ where \ get_Mcause_Cause v = ( (subrange_vec_dec(Mcause_bits v) (( 30 :: int)::ii) (( 0 :: int)::ii) :: 31 Word.word))\ for v :: " Mcause " \ \\val _set_Mcause_Cause : register_ref regstate register_value Mcause -> mword ty31 -> M unit\\ definition set_Mcause_Cause :: \((regstate),(register_value),(Mcause))register_ref \(31)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mcause_Cause r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mcause_bits := ((update_subrange_vec_dec(Mcause_bits r) (( 30 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mcause))register_ref " and v :: "(31)Word.word " \ \\val _update_Mcause_Cause : Mcause -> mword ty31 -> Mcause\\ definition update_Mcause_Cause :: \ Mcause \(31)Word.word \ Mcause \ where \ update_Mcause_Cause v x = ( ( v (| Mcause_bits := ((update_subrange_vec_dec(Mcause_bits v) (( 30 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mcause " and x :: "(31)Word.word " \ \\val _get_Mcause_IsInterrupt : Mcause -> mword ty1\\ definition get_Mcause_IsInterrupt :: \ Mcause \(1)Word.word \ where \ get_Mcause_IsInterrupt v = ( (subrange_vec_dec(Mcause_bits v) (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word))\ for v :: " Mcause " \ \\val _set_Mcause_IsInterrupt : register_ref regstate register_value Mcause -> mword ty1 -> M unit\\ definition set_Mcause_IsInterrupt :: \((regstate),(register_value),(Mcause))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Mcause_IsInterrupt r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Mcause_bits := ((update_subrange_vec_dec(Mcause_bits r) (( 31 :: int)::ii) (( 31 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Mcause))register_ref " and v :: "(1)Word.word " \ \\val _update_Mcause_IsInterrupt : Mcause -> mword ty1 -> Mcause\\ definition update_Mcause_IsInterrupt :: \ Mcause \(1)Word.word \ Mcause \ where \ update_Mcause_IsInterrupt v x = ( ( v (| Mcause_bits := ((update_subrange_vec_dec(Mcause_bits v) (( 31 :: int)::ii) (( 31 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Mcause " and x :: "(1)Word.word " \ \\val tvec_addr : Mtvec -> Mcause -> maybe (mword ty32)\\ definition tvec_addr :: \ Mtvec \ Mcause \((32)Word.word)option \ where \ tvec_addr (m :: Mtvec) (c :: Mcause) = ( (let (base :: xlenbits) = ((concat_vec ((get_Mtvec_Base m :: 30 Word.word)) ( 0b00 :: 2 Word.word) :: 32 Word.word)) in (case ((trapVectorMode_of_bits ((get_Mtvec_Mode m :: 2 Word.word)))) of TV_Direct => Some base | TV_Vector => if (((((get_Mcause_IsInterrupt c :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some ((add_vec base ((shiftl ((EXTZ (( 32 :: int)::ii) ((get_Mcause_Cause c :: 31 Word.word)) :: 32 Word.word)) (( 2 :: int)::ii) :: 32 Word.word)) :: 32 Word.word)) else Some base | TV_Reserved => None )))\ for m :: " Mtvec " and c :: " Mcause " \ \\val legalize_xepc : mword ty32 -> M (mword ty32)\\ definition legalize_xepc :: \(32)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ legalize_xepc v = ( or_boolM (return (((((sys_enable_writable_misa () )) \ ((sys_enable_rvc () )))))) (read_reg misa_ref \ ((\ (w__0 :: Misa) . return (((((get_Misa_C w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))) \ ((\ (w__1 :: bool) . return (if w__1 then (update_vec_dec v (( 0 :: int)::ii) B0 :: 32 Word.word) else (and_vec v ((EXTS (( 32 :: int)::ii) ( 0b100 :: 3 Word.word) :: 32 Word.word)) :: 32 Word.word)))))\ for v :: "(32)Word.word " \ \\val pc_alignment_mask : unit -> M (mword ty32)\\ definition pc_alignment_mask :: \ unit \((register_value),((32)Word.word),(exception))monad \ where \ pc_alignment_mask _ = ( read_reg misa_ref \ ((\ (w__0 :: Misa) . return ((not_vec ((EXTZ (( 32 :: int)::ii) (if (((((get_Misa_C w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then ( 0b00 :: 2 Word.word) else ( 0b10 :: 2 Word.word)) :: 32 Word.word)) :: 32 Word.word)))))\ \ \\val undefined_Counteren : unit -> M Counteren\\ definition undefined_Counteren :: \ unit \((register_value),(Counteren),(exception))monad \ where \ undefined_Counteren _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Counteren_bits = w__0 |)))))\ \ \\val Mk_Counteren : mword ty32 -> Counteren\\ definition Mk_Counteren :: \(32)Word.word \ Counteren \ where \ Mk_Counteren v = ( (| Counteren_bits = v |) )\ for v :: "(32)Word.word " definition get_Counteren_bits :: \ Counteren \(32)Word.word \ where \ get_Counteren_bits v = ( (subrange_vec_dec(Counteren_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Counteren " definition set_Counteren_bits :: \((regstate),(register_value),(Counteren))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Counteren_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Counteren))register_ref " and v :: "(32)Word.word " definition update_Counteren_bits :: \ Counteren \(32)Word.word \ Counteren \ where \ update_Counteren_bits v x = ( ( v (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Counteren " and x :: "(32)Word.word " \ \\val _get_Counteren_CY : Counteren -> mword ty1\\ definition get_Counteren_CY :: \ Counteren \(1)Word.word \ where \ get_Counteren_CY v = ( (subrange_vec_dec(Counteren_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Counteren " \ \\val _set_Counteren_CY : register_ref regstate register_value Counteren -> mword ty1 -> M unit\\ definition set_Counteren_CY :: \((regstate),(register_value),(Counteren))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Counteren_CY r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Counteren))register_ref " and v :: "(1)Word.word " \ \\val _update_Counteren_CY : Counteren -> mword ty1 -> Counteren\\ definition update_Counteren_CY :: \ Counteren \(1)Word.word \ Counteren \ where \ update_Counteren_CY v x = ( ( v (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Counteren " and x :: "(1)Word.word " \ \\val _update_Counterin_CY : Counterin -> mword ty1 -> Counterin\\ \ \\val _get_Counterin_CY : Counterin -> mword ty1\\ \ \\val _set_Counterin_CY : register_ref regstate register_value Counterin -> mword ty1 -> M unit\\ \ \\val _get_Counteren_HPM : Counteren -> mword ty29\\ definition get_Counteren_HPM :: \ Counteren \(29)Word.word \ where \ get_Counteren_HPM v = ( (subrange_vec_dec(Counteren_bits v) (( 31 :: int)::ii) (( 3 :: int)::ii) :: 29 Word.word))\ for v :: " Counteren " \ \\val _set_Counteren_HPM : register_ref regstate register_value Counteren -> mword ty29 -> M unit\\ definition set_Counteren_HPM :: \((regstate),(register_value),(Counteren))register_ref \(29)Word.word \((register_value),(unit),(exception))monad \ where \ set_Counteren_HPM r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits r) (( 31 :: int)::ii) (( 3 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Counteren))register_ref " and v :: "(29)Word.word " \ \\val _update_Counteren_HPM : Counteren -> mword ty29 -> Counteren\\ definition update_Counteren_HPM :: \ Counteren \(29)Word.word \ Counteren \ where \ update_Counteren_HPM v x = ( ( v (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits v) (( 31 :: int)::ii) (( 3 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Counteren " and x :: "(29)Word.word " \ \\val _get_Counteren_IR : Counteren -> mword ty1\\ definition get_Counteren_IR :: \ Counteren \(1)Word.word \ where \ get_Counteren_IR v = ( (subrange_vec_dec(Counteren_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word))\ for v :: " Counteren " \ \\val _set_Counteren_IR : register_ref regstate register_value Counteren -> mword ty1 -> M unit\\ definition set_Counteren_IR :: \((regstate),(register_value),(Counteren))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Counteren_IR r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits r) (( 2 :: int)::ii) (( 2 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Counteren))register_ref " and v :: "(1)Word.word " \ \\val _update_Counteren_IR : Counteren -> mword ty1 -> Counteren\\ definition update_Counteren_IR :: \ Counteren \(1)Word.word \ Counteren \ where \ update_Counteren_IR v x = ( ( v (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Counteren " and x :: "(1)Word.word " \ \\val _update_Counterin_IR : Counterin -> mword ty1 -> Counterin\\ \ \\val _get_Counterin_IR : Counterin -> mword ty1\\ \ \\val _set_Counterin_IR : register_ref regstate register_value Counterin -> mword ty1 -> M unit\\ \ \\val _get_Counteren_TM : Counteren -> mword ty1\\ definition get_Counteren_TM :: \ Counteren \(1)Word.word \ where \ get_Counteren_TM v = ( (subrange_vec_dec(Counteren_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Counteren " \ \\val _set_Counteren_TM : register_ref regstate register_value Counteren -> mword ty1 -> M unit\\ definition set_Counteren_TM :: \((regstate),(register_value),(Counteren))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Counteren_TM r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Counteren))register_ref " and v :: "(1)Word.word " \ \\val _update_Counteren_TM : Counteren -> mword ty1 -> Counteren\\ definition update_Counteren_TM :: \ Counteren \(1)Word.word \ Counteren \ where \ update_Counteren_TM v x = ( ( v (| Counteren_bits := ((update_subrange_vec_dec(Counteren_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Counteren " and x :: "(1)Word.word " \ \\val legalize_mcounteren : Counteren -> mword ty32 -> Counteren\\ definition legalize_mcounteren :: \ Counteren \(32)Word.word \ Counteren \ where \ legalize_mcounteren (c :: Counteren) (v :: xlenbits) = ( (let c = (update_Counteren_IR c (vec_of_bits [access_vec_dec v (( 2 :: int)::ii)] :: 1 Word.word)) in (let c = (update_Counteren_TM c (vec_of_bits [access_vec_dec v (( 1 :: int)::ii)] :: 1 Word.word)) in update_Counteren_CY c (vec_of_bits [access_vec_dec v (( 0 :: int)::ii)] :: 1 Word.word))))\ for c :: " Counteren " and v :: "(32)Word.word " \ \\val legalize_scounteren : Counteren -> mword ty32 -> Counteren\\ definition legalize_scounteren :: \ Counteren \(32)Word.word \ Counteren \ where \ legalize_scounteren (c :: Counteren) (v :: xlenbits) = ( (let c = (update_Counteren_IR c (vec_of_bits [access_vec_dec v (( 2 :: int)::ii)] :: 1 Word.word)) in (let c = (update_Counteren_TM c (vec_of_bits [access_vec_dec v (( 1 :: int)::ii)] :: 1 Word.word)) in update_Counteren_CY c (vec_of_bits [access_vec_dec v (( 0 :: int)::ii)] :: 1 Word.word))))\ for c :: " Counteren " and v :: "(32)Word.word " \ \\val undefined_Counterin : unit -> M Counterin\\ definition undefined_Counterin :: \ unit \((register_value),(Counterin),(exception))monad \ where \ undefined_Counterin _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Counterin_bits = w__0 |)))))\ \ \\val Mk_Counterin : mword ty32 -> Counterin\\ definition Mk_Counterin :: \(32)Word.word \ Counterin \ where \ Mk_Counterin v = ( (| Counterin_bits = v |) )\ for v :: "(32)Word.word " definition get_Counterin_bits :: \ Counterin \(32)Word.word \ where \ get_Counterin_bits v = ( (subrange_vec_dec(Counterin_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Counterin " definition set_Counterin_bits :: \((regstate),(register_value),(Counterin))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Counterin_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Counterin_bits := ((update_subrange_vec_dec(Counterin_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Counterin))register_ref " and v :: "(32)Word.word " definition update_Counterin_bits :: \ Counterin \(32)Word.word \ Counterin \ where \ update_Counterin_bits v x = ( ( v (| Counterin_bits := ((update_subrange_vec_dec(Counterin_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Counterin " and x :: "(32)Word.word " definition get_Counterin_CY :: \ Counterin \(1)Word.word \ where \ get_Counterin_CY v = ( (subrange_vec_dec(Counterin_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Counterin " definition set_Counterin_CY :: \((regstate),(register_value),(Counterin))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Counterin_CY r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Counterin_bits := ((update_subrange_vec_dec(Counterin_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Counterin))register_ref " and v :: "(1)Word.word " definition update_Counterin_CY :: \ Counterin \(1)Word.word \ Counterin \ where \ update_Counterin_CY v x = ( ( v (| Counterin_bits := ((update_subrange_vec_dec(Counterin_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Counterin " and x :: "(1)Word.word " definition get_Counterin_IR :: \ Counterin \(1)Word.word \ where \ get_Counterin_IR v = ( (subrange_vec_dec(Counterin_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word))\ for v :: " Counterin " definition set_Counterin_IR :: \((regstate),(register_value),(Counterin))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Counterin_IR r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Counterin_bits := ((update_subrange_vec_dec(Counterin_bits r) (( 2 :: int)::ii) (( 2 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Counterin))register_ref " and v :: "(1)Word.word " definition update_Counterin_IR :: \ Counterin \(1)Word.word \ Counterin \ where \ update_Counterin_IR v x = ( ( v (| Counterin_bits := ((update_subrange_vec_dec(Counterin_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Counterin " and x :: "(1)Word.word " \ \\val legalize_mcountinhibit : Counterin -> mword ty32 -> Counterin\\ definition legalize_mcountinhibit :: \ Counterin \(32)Word.word \ Counterin \ where \ legalize_mcountinhibit (c :: Counterin) (v :: xlenbits) = ( (let c = (update_Counterin_IR c (vec_of_bits [access_vec_dec v (( 2 :: int)::ii)] :: 1 Word.word)) in update_Counterin_CY c (vec_of_bits [access_vec_dec v (( 0 :: int)::ii)] :: 1 Word.word)))\ for c :: " Counterin " and v :: "(32)Word.word " \ \\val retire_instruction : unit -> M unit\\ definition retire_instruction :: \ unit \((register_value),(unit),(exception))monad \ where \ retire_instruction _ = ( read_reg minstret_written_ref \ ((\ (w__0 :: bool) . if (((w__0 = True))) then write_reg minstret_written_ref False else read_reg mcountinhibit_ref \ ((\ (w__1 :: Counterin) . if (((((get_Counterin_IR w__1 :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then (read_reg minstret_ref :: ( 64 Word.word) M) \ ((\ (w__2 :: 64 Word.word) . write_reg minstret_ref ((add_vec_int w__2 (( 1 :: int)::ii) :: 64 Word.word)))) else return () )))))\ \ \\val undefined_Sstatus : unit -> M Sstatus\\ definition undefined_Sstatus :: \ unit \((register_value),(Sstatus),(exception))monad \ where \ undefined_Sstatus _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Sstatus_bits = w__0 |)))))\ \ \\val Mk_Sstatus : mword ty32 -> Sstatus\\ definition Mk_Sstatus :: \(32)Word.word \ Sstatus \ where \ Mk_Sstatus v = ( (| Sstatus_bits = v |) )\ for v :: "(32)Word.word " definition get_Sstatus_bits :: \ Sstatus \(32)Word.word \ where \ get_Sstatus_bits v = ( (subrange_vec_dec(Sstatus_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Sstatus " definition set_Sstatus_bits :: \((regstate),(register_value),(Sstatus))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(32)Word.word " definition update_Sstatus_bits :: \ Sstatus \(32)Word.word \ Sstatus \ where \ update_Sstatus_bits v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(32)Word.word " definition get_Sstatus_FS :: \ Sstatus \(2)Word.word \ where \ get_Sstatus_FS v = ( (subrange_vec_dec(Sstatus_bits v) (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word))\ for v :: " Sstatus " definition set_Sstatus_FS :: \((regstate),(register_value),(Sstatus))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_FS r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 14 :: int)::ii) (( 13 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(2)Word.word " definition update_Sstatus_FS :: \ Sstatus \(2)Word.word \ Sstatus \ where \ update_Sstatus_FS v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 14 :: int)::ii) (( 13 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(2)Word.word " definition get_Sstatus_MXR :: \ Sstatus \(1)Word.word \ where \ get_Sstatus_MXR v = ( (subrange_vec_dec(Sstatus_bits v) (( 19 :: int)::ii) (( 19 :: int)::ii) :: 1 Word.word))\ for v :: " Sstatus " definition set_Sstatus_MXR :: \((regstate),(register_value),(Sstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_MXR r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 19 :: int)::ii) (( 19 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(1)Word.word " definition update_Sstatus_MXR :: \ Sstatus \(1)Word.word \ Sstatus \ where \ update_Sstatus_MXR v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 19 :: int)::ii) (( 19 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(1)Word.word " definition get_Sstatus_SD :: \ Sstatus \(1)Word.word \ where \ get_Sstatus_SD v = ( (subrange_vec_dec(Sstatus_bits v) (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word))\ for v :: " Sstatus " definition set_Sstatus_SD :: \((regstate),(register_value),(Sstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_SD r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 31 :: int)::ii) (( 31 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(1)Word.word " definition update_Sstatus_SD :: \ Sstatus \(1)Word.word \ Sstatus \ where \ update_Sstatus_SD v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 31 :: int)::ii) (( 31 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(1)Word.word " definition get_Sstatus_SIE :: \ Sstatus \(1)Word.word \ where \ get_Sstatus_SIE v = ( (subrange_vec_dec(Sstatus_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Sstatus " definition set_Sstatus_SIE :: \((regstate),(register_value),(Sstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_SIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(1)Word.word " definition update_Sstatus_SIE :: \ Sstatus \(1)Word.word \ Sstatus \ where \ update_Sstatus_SIE v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(1)Word.word " definition get_Sstatus_SPIE :: \ Sstatus \(1)Word.word \ where \ get_Sstatus_SPIE v = ( (subrange_vec_dec(Sstatus_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " Sstatus " definition set_Sstatus_SPIE :: \((regstate),(register_value),(Sstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_SPIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(1)Word.word " definition update_Sstatus_SPIE :: \ Sstatus \(1)Word.word \ Sstatus \ where \ update_Sstatus_SPIE v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(1)Word.word " definition get_Sstatus_SPP :: \ Sstatus \(1)Word.word \ where \ get_Sstatus_SPP v = ( (subrange_vec_dec(Sstatus_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word))\ for v :: " Sstatus " definition set_Sstatus_SPP :: \((regstate),(register_value),(Sstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_SPP r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 8 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(1)Word.word " definition update_Sstatus_SPP :: \ Sstatus \(1)Word.word \ Sstatus \ where \ update_Sstatus_SPP v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(1)Word.word " definition get_Sstatus_SUM :: \ Sstatus \(1)Word.word \ where \ get_Sstatus_SUM v = ( (subrange_vec_dec(Sstatus_bits v) (( 18 :: int)::ii) (( 18 :: int)::ii) :: 1 Word.word))\ for v :: " Sstatus " definition set_Sstatus_SUM :: \((regstate),(register_value),(Sstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_SUM r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 18 :: int)::ii) (( 18 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(1)Word.word " definition update_Sstatus_SUM :: \ Sstatus \(1)Word.word \ Sstatus \ where \ update_Sstatus_SUM v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 18 :: int)::ii) (( 18 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(1)Word.word " definition get_Sstatus_UIE :: \ Sstatus \(1)Word.word \ where \ get_Sstatus_UIE v = ( (subrange_vec_dec(Sstatus_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Sstatus " definition set_Sstatus_UIE :: \((regstate),(register_value),(Sstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_UIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(1)Word.word " definition update_Sstatus_UIE :: \ Sstatus \(1)Word.word \ Sstatus \ where \ update_Sstatus_UIE v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(1)Word.word " definition get_Sstatus_UPIE :: \ Sstatus \(1)Word.word \ where \ get_Sstatus_UPIE v = ( (subrange_vec_dec(Sstatus_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Sstatus " definition set_Sstatus_UPIE :: \((regstate),(register_value),(Sstatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_UPIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(1)Word.word " definition update_Sstatus_UPIE :: \ Sstatus \(1)Word.word \ Sstatus \ where \ update_Sstatus_UPIE v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(1)Word.word " definition get_Sstatus_XS :: \ Sstatus \(2)Word.word \ where \ get_Sstatus_XS v = ( (subrange_vec_dec(Sstatus_bits v) (( 16 :: int)::ii) (( 15 :: int)::ii) :: 2 Word.word))\ for v :: " Sstatus " definition set_Sstatus_XS :: \((regstate),(register_value),(Sstatus))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sstatus_XS r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits r) (( 16 :: int)::ii) (( 15 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sstatus))register_ref " and v :: "(2)Word.word " definition update_Sstatus_XS :: \ Sstatus \(2)Word.word \ Sstatus \ where \ update_Sstatus_XS v x = ( ( v (| Sstatus_bits := ((update_subrange_vec_dec(Sstatus_bits v) (( 16 :: int)::ii) (( 15 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sstatus " and x :: "(2)Word.word " \ \\val get_sstatus_UXL : Sstatus -> mword ty2\\ definition get_sstatus_UXL :: \ Sstatus \(2)Word.word \ where \ get_sstatus_UXL s = ( (let m = (Mk_Mstatus ((get_Sstatus_bits s :: 32 Word.word))) in (get_mstatus_UXL m :: 2 Word.word)))\ for s :: " Sstatus " \ \\val set_sstatus_UXL : Sstatus -> mword ty2 -> Sstatus\\ definition set_sstatus_UXL :: \ Sstatus \(2)Word.word \ Sstatus \ where \ set_sstatus_UXL (s :: Sstatus) (a :: arch_xlen) = ( (let m = (Mk_Mstatus ((get_Sstatus_bits s :: 32 Word.word))) in (let m = (set_mstatus_UXL m a) in Mk_Sstatus ((get_Mstatus_bits m :: 32 Word.word)))))\ for s :: " Sstatus " and a :: "(2)Word.word " \ \\val lower_mstatus : Mstatus -> Sstatus\\ definition lower_mstatus :: \ Mstatus \ Sstatus \ where \ lower_mstatus m = ( (let s = (Mk_Sstatus ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) in (let s = (update_Sstatus_SD s ((get_Mstatus_SD m :: 1 Word.word))) in (let s = (set_sstatus_UXL s ((get_mstatus_UXL m :: 2 Word.word))) in (let s = (update_Sstatus_MXR s ((get_Mstatus_MXR m :: 1 Word.word))) in (let s = (update_Sstatus_SUM s ((get_Mstatus_SUM m :: 1 Word.word))) in (let s = (update_Sstatus_XS s ((get_Mstatus_XS m :: 2 Word.word))) in (let s = (update_Sstatus_FS s ((get_Mstatus_FS m :: 2 Word.word))) in (let s = (update_Sstatus_SPP s ((get_Mstatus_SPP m :: 1 Word.word))) in (let s = (update_Sstatus_SPIE s ((get_Mstatus_SPIE m :: 1 Word.word))) in (let s = (update_Sstatus_UPIE s ((get_Mstatus_UPIE m :: 1 Word.word))) in (let s = (update_Sstatus_SIE s ((get_Mstatus_SIE m :: 1 Word.word))) in update_Sstatus_UIE s ((get_Mstatus_UIE m :: 1 Word.word))))))))))))))\ for m :: " Mstatus " \ \\val lift_sstatus : Mstatus -> Sstatus -> M Mstatus\\ definition lift_sstatus :: \ Mstatus \ Sstatus \((register_value),(Mstatus),(exception))monad \ where \ lift_sstatus (m :: Mstatus) (s :: Sstatus) = ( (let m = (update_Mstatus_MXR m ((get_Sstatus_MXR s :: 1 Word.word))) in (let m = (update_Mstatus_SUM m ((get_Sstatus_SUM s :: 1 Word.word))) in (let m = (update_Mstatus_XS m ((get_Sstatus_XS s :: 2 Word.word))) in (let m = (update_Mstatus_FS m ((get_Sstatus_FS s :: 2 Word.word))) in or_boolM (extStatus_of_bits ((get_Mstatus_FS m :: 2 Word.word)) \ ((\ (w__0 :: ExtStatus) . return (((w__0 = Dirty)))))) (extStatus_of_bits ((get_Mstatus_XS m :: 2 Word.word)) \ ((\ (w__1 :: ExtStatus) . return (((w__1 = Dirty)))))) \ ((\ dirty . (let m = (update_Mstatus_SD m ((bool_to_bits dirty :: 1 Word.word))) in (let m = (update_Mstatus_SPP m ((get_Sstatus_SPP s :: 1 Word.word))) in (let m = (update_Mstatus_SPIE m ((get_Sstatus_SPIE s :: 1 Word.word))) in (let m = (update_Mstatus_UPIE m ((get_Sstatus_UPIE s :: 1 Word.word))) in (let m = (update_Mstatus_SIE m ((get_Sstatus_SIE s :: 1 Word.word))) in (let m = (update_Mstatus_UIE m ((get_Sstatus_UIE s :: 1 Word.word))) in return m)))))))))))))\ for m :: " Mstatus " and s :: " Sstatus " \ \\val legalize_sstatus : Mstatus -> mword ty32 -> M Mstatus\\ definition legalize_sstatus :: \ Mstatus \(32)Word.word \((register_value),(Mstatus),(exception))monad \ where \ legalize_sstatus (m :: Mstatus) (v :: xlenbits) = ( lift_sstatus m ((Mk_Sstatus v)) \ ((\ (w__0 :: Mstatus) . legalize_mstatus m ((get_Mstatus_bits w__0 :: 32 Word.word)))))\ for m :: " Mstatus " and v :: "(32)Word.word " \ \\val undefined_Sedeleg : unit -> M Sedeleg\\ definition undefined_Sedeleg :: \ unit \((register_value),(Sedeleg),(exception))monad \ where \ undefined_Sedeleg _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Sedeleg_bits = w__0 |)))))\ \ \\val Mk_Sedeleg : mword ty32 -> Sedeleg\\ definition Mk_Sedeleg :: \(32)Word.word \ Sedeleg \ where \ Mk_Sedeleg v = ( (| Sedeleg_bits = v |) )\ for v :: "(32)Word.word " definition get_Sedeleg_bits :: \ Sedeleg \(32)Word.word \ where \ get_Sedeleg_bits v = ( (subrange_vec_dec(Sedeleg_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_bits :: \((regstate),(register_value),(Sedeleg))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(32)Word.word " definition update_Sedeleg_bits :: \ Sedeleg \(32)Word.word \ Sedeleg \ where \ update_Sedeleg_bits v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(32)Word.word " definition get_Sedeleg_Breakpoint :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_Breakpoint v = ( (subrange_vec_dec(Sedeleg_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_Breakpoint :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_Breakpoint r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 3 :: int)::ii) (( 3 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_Breakpoint :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_Breakpoint v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " definition get_Sedeleg_Fetch_Access_Fault :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_Fetch_Access_Fault v = ( (subrange_vec_dec(Sedeleg_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_Fetch_Access_Fault :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_Fetch_Access_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_Fetch_Access_Fault :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_Fetch_Access_Fault v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " definition get_Sedeleg_Fetch_Addr_Align :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_Fetch_Addr_Align v = ( (subrange_vec_dec(Sedeleg_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_Fetch_Addr_Align :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_Fetch_Addr_Align r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_Fetch_Addr_Align :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_Fetch_Addr_Align v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " definition get_Sedeleg_Illegal_Instr :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_Illegal_Instr v = ( (subrange_vec_dec(Sedeleg_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_Illegal_Instr :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_Illegal_Instr r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 2 :: int)::ii) (( 2 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_Illegal_Instr :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_Illegal_Instr v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " definition get_Sedeleg_Load_Access_Fault :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_Load_Access_Fault v = ( (subrange_vec_dec(Sedeleg_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_Load_Access_Fault :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_Load_Access_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_Load_Access_Fault :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_Load_Access_Fault v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " definition get_Sedeleg_Load_Addr_Align :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_Load_Addr_Align v = ( (subrange_vec_dec(Sedeleg_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_Load_Addr_Align :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_Load_Addr_Align r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_Load_Addr_Align :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_Load_Addr_Align v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " definition get_Sedeleg_SAMO_Access_Fault :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_SAMO_Access_Fault v = ( (subrange_vec_dec(Sedeleg_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_SAMO_Access_Fault :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_SAMO_Access_Fault r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 7 :: int)::ii) (( 7 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_SAMO_Access_Fault :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_SAMO_Access_Fault v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " definition get_Sedeleg_SAMO_Addr_Align :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_SAMO_Addr_Align v = ( (subrange_vec_dec(Sedeleg_bits v) (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_SAMO_Addr_Align :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_SAMO_Addr_Align r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 6 :: int)::ii) (( 6 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_SAMO_Addr_Align :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_SAMO_Addr_Align v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 6 :: int)::ii) (( 6 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " definition get_Sedeleg_UEnvCall :: \ Sedeleg \(1)Word.word \ where \ get_Sedeleg_UEnvCall v = ( (subrange_vec_dec(Sedeleg_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word))\ for v :: " Sedeleg " definition set_Sedeleg_UEnvCall :: \((regstate),(register_value),(Sedeleg))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sedeleg_UEnvCall r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits r) (( 8 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sedeleg))register_ref " and v :: "(1)Word.word " definition update_Sedeleg_UEnvCall :: \ Sedeleg \(1)Word.word \ Sedeleg \ where \ update_Sedeleg_UEnvCall v x = ( ( v (| Sedeleg_bits := ((update_subrange_vec_dec(Sedeleg_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sedeleg " and x :: "(1)Word.word " \ \\val legalize_sedeleg : Sedeleg -> mword ty32 -> Sedeleg\\ definition legalize_sedeleg :: \ Sedeleg \(32)Word.word \ Sedeleg \ where \ legalize_sedeleg (s :: Sedeleg) (v :: xlenbits) = ( Mk_Sedeleg ((EXTZ (( 32 :: int)::ii) ((subrange_vec_dec v (( 8 :: int)::ii) (( 0 :: int)::ii) :: 9 Word.word)) :: 32 Word.word)))\ for s :: " Sedeleg " and v :: "(32)Word.word " \ \\val undefined_Sinterrupts : unit -> M Sinterrupts\\ definition undefined_Sinterrupts :: \ unit \((register_value),(Sinterrupts),(exception))monad \ where \ undefined_Sinterrupts _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Sinterrupts_bits = w__0 |)))))\ \ \\val Mk_Sinterrupts : mword ty32 -> Sinterrupts\\ definition Mk_Sinterrupts :: \(32)Word.word \ Sinterrupts \ where \ Mk_Sinterrupts v = ( (| Sinterrupts_bits = v |) )\ for v :: "(32)Word.word " definition get_Sinterrupts_bits :: \ Sinterrupts \(32)Word.word \ where \ get_Sinterrupts_bits v = ( (subrange_vec_dec(Sinterrupts_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Sinterrupts " definition set_Sinterrupts_bits :: \((regstate),(register_value),(Sinterrupts))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sinterrupts_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sinterrupts))register_ref " and v :: "(32)Word.word " definition update_Sinterrupts_bits :: \ Sinterrupts \(32)Word.word \ Sinterrupts \ where \ update_Sinterrupts_bits v x = ( ( v (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sinterrupts " and x :: "(32)Word.word " definition get_Sinterrupts_SEI :: \ Sinterrupts \(1)Word.word \ where \ get_Sinterrupts_SEI v = ( (subrange_vec_dec(Sinterrupts_bits v) (( 9 :: int)::ii) (( 9 :: int)::ii) :: 1 Word.word))\ for v :: " Sinterrupts " definition set_Sinterrupts_SEI :: \((regstate),(register_value),(Sinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sinterrupts_SEI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits r) (( 9 :: int)::ii) (( 9 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sinterrupts))register_ref " and v :: "(1)Word.word " definition update_Sinterrupts_SEI :: \ Sinterrupts \(1)Word.word \ Sinterrupts \ where \ update_Sinterrupts_SEI v x = ( ( v (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits v) (( 9 :: int)::ii) (( 9 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sinterrupts " and x :: "(1)Word.word " definition get_Sinterrupts_SSI :: \ Sinterrupts \(1)Word.word \ where \ get_Sinterrupts_SSI v = ( (subrange_vec_dec(Sinterrupts_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Sinterrupts " definition set_Sinterrupts_SSI :: \((regstate),(register_value),(Sinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sinterrupts_SSI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sinterrupts))register_ref " and v :: "(1)Word.word " definition update_Sinterrupts_SSI :: \ Sinterrupts \(1)Word.word \ Sinterrupts \ where \ update_Sinterrupts_SSI v x = ( ( v (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sinterrupts " and x :: "(1)Word.word " definition get_Sinterrupts_STI :: \ Sinterrupts \(1)Word.word \ where \ get_Sinterrupts_STI v = ( (subrange_vec_dec(Sinterrupts_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " Sinterrupts " definition set_Sinterrupts_STI :: \((regstate),(register_value),(Sinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sinterrupts_STI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sinterrupts))register_ref " and v :: "(1)Word.word " definition update_Sinterrupts_STI :: \ Sinterrupts \(1)Word.word \ Sinterrupts \ where \ update_Sinterrupts_STI v x = ( ( v (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sinterrupts " and x :: "(1)Word.word " definition get_Sinterrupts_UEI :: \ Sinterrupts \(1)Word.word \ where \ get_Sinterrupts_UEI v = ( (subrange_vec_dec(Sinterrupts_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word))\ for v :: " Sinterrupts " definition set_Sinterrupts_UEI :: \((regstate),(register_value),(Sinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sinterrupts_UEI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits r) (( 8 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sinterrupts))register_ref " and v :: "(1)Word.word " definition update_Sinterrupts_UEI :: \ Sinterrupts \(1)Word.word \ Sinterrupts \ where \ update_Sinterrupts_UEI v x = ( ( v (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sinterrupts " and x :: "(1)Word.word " definition get_Sinterrupts_USI :: \ Sinterrupts \(1)Word.word \ where \ get_Sinterrupts_USI v = ( (subrange_vec_dec(Sinterrupts_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Sinterrupts " definition set_Sinterrupts_USI :: \((regstate),(register_value),(Sinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sinterrupts_USI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sinterrupts))register_ref " and v :: "(1)Word.word " definition update_Sinterrupts_USI :: \ Sinterrupts \(1)Word.word \ Sinterrupts \ where \ update_Sinterrupts_USI v x = ( ( v (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sinterrupts " and x :: "(1)Word.word " definition get_Sinterrupts_UTI :: \ Sinterrupts \(1)Word.word \ where \ get_Sinterrupts_UTI v = ( (subrange_vec_dec(Sinterrupts_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Sinterrupts " definition set_Sinterrupts_UTI :: \((regstate),(register_value),(Sinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Sinterrupts_UTI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Sinterrupts))register_ref " and v :: "(1)Word.word " definition update_Sinterrupts_UTI :: \ Sinterrupts \(1)Word.word \ Sinterrupts \ where \ update_Sinterrupts_UTI v x = ( ( v (| Sinterrupts_bits := ((update_subrange_vec_dec(Sinterrupts_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Sinterrupts " and x :: "(1)Word.word " \ \\val lower_mip : Minterrupts -> Minterrupts -> Sinterrupts\\ definition lower_mip :: \ Minterrupts \ Minterrupts \ Sinterrupts \ where \ lower_mip (m :: Minterrupts) (d :: Minterrupts) = ( (let (s :: Sinterrupts) = (Mk_Sinterrupts ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) in (let s = (update_Sinterrupts_SEI s ((and_vec ((get_Minterrupts_SEI m :: 1 Word.word)) ((get_Minterrupts_SEI d :: 1 Word.word)) :: 1 Word.word))) in (let s = (update_Sinterrupts_STI s ((and_vec ((get_Minterrupts_STI m :: 1 Word.word)) ((get_Minterrupts_STI d :: 1 Word.word)) :: 1 Word.word))) in (let s = (update_Sinterrupts_SSI s ((and_vec ((get_Minterrupts_SSI m :: 1 Word.word)) ((get_Minterrupts_SSI d :: 1 Word.word)) :: 1 Word.word))) in (let s = (update_Sinterrupts_UEI s ((and_vec ((get_Minterrupts_UEI m :: 1 Word.word)) ((get_Minterrupts_UEI d :: 1 Word.word)) :: 1 Word.word))) in (let s = (update_Sinterrupts_UTI s ((and_vec ((get_Minterrupts_UTI m :: 1 Word.word)) ((get_Minterrupts_UTI d :: 1 Word.word)) :: 1 Word.word))) in update_Sinterrupts_USI s ((and_vec ((get_Minterrupts_USI m :: 1 Word.word)) ((get_Minterrupts_USI d :: 1 Word.word)) :: 1 Word.word)))))))))\ for m :: " Minterrupts " and d :: " Minterrupts " \ \\val lower_mie : Minterrupts -> Minterrupts -> Sinterrupts\\ definition lower_mie :: \ Minterrupts \ Minterrupts \ Sinterrupts \ where \ lower_mie (m :: Minterrupts) (d :: Minterrupts) = ( (let (s :: Sinterrupts) = (Mk_Sinterrupts ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) in (let s = (update_Sinterrupts_SEI s ((and_vec ((get_Minterrupts_SEI m :: 1 Word.word)) ((get_Minterrupts_SEI d :: 1 Word.word)) :: 1 Word.word))) in (let s = (update_Sinterrupts_STI s ((and_vec ((get_Minterrupts_STI m :: 1 Word.word)) ((get_Minterrupts_STI d :: 1 Word.word)) :: 1 Word.word))) in (let s = (update_Sinterrupts_SSI s ((and_vec ((get_Minterrupts_SSI m :: 1 Word.word)) ((get_Minterrupts_SSI d :: 1 Word.word)) :: 1 Word.word))) in (let s = (update_Sinterrupts_UEI s ((and_vec ((get_Minterrupts_UEI m :: 1 Word.word)) ((get_Minterrupts_UEI d :: 1 Word.word)) :: 1 Word.word))) in (let s = (update_Sinterrupts_UTI s ((and_vec ((get_Minterrupts_UTI m :: 1 Word.word)) ((get_Minterrupts_UTI d :: 1 Word.word)) :: 1 Word.word))) in update_Sinterrupts_USI s ((and_vec ((get_Minterrupts_USI m :: 1 Word.word)) ((get_Minterrupts_USI d :: 1 Word.word)) :: 1 Word.word)))))))))\ for m :: " Minterrupts " and d :: " Minterrupts " \ \\val lift_sip : Minterrupts -> Minterrupts -> Sinterrupts -> M Minterrupts\\ definition lift_sip :: \ Minterrupts \ Minterrupts \ Sinterrupts \((register_value),(Minterrupts),(exception))monad \ where \ lift_sip (o1 :: Minterrupts) (d :: Minterrupts) (s :: Sinterrupts) = ( (let (m :: Minterrupts) = o1 in (let m = (if (((((get_Minterrupts_SSI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_SSI m ((get_Sinterrupts_SSI s :: 1 Word.word)) else m) in haveNExt () \ ((\ (w__0 :: bool) . return (if w__0 then (let m = (if (((((get_Minterrupts_UEI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_UEI m ((get_Sinterrupts_UEI s :: 1 Word.word)) else m) in if (((((get_Minterrupts_USI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_USI m ((get_Sinterrupts_USI s :: 1 Word.word)) else m) else m))))))\ for o1 :: " Minterrupts " and d :: " Minterrupts " and s :: " Sinterrupts " \ \\val legalize_sip : Minterrupts -> Minterrupts -> mword ty32 -> M Minterrupts\\ definition legalize_sip :: \ Minterrupts \ Minterrupts \(32)Word.word \((register_value),(Minterrupts),(exception))monad \ where \ legalize_sip (m :: Minterrupts) (d :: Minterrupts) (v :: xlenbits) = ( lift_sip m d ((Mk_Sinterrupts v)))\ for m :: " Minterrupts " and d :: " Minterrupts " and v :: "(32)Word.word " \ \\val lift_sie : Minterrupts -> Minterrupts -> Sinterrupts -> M Minterrupts\\ definition lift_sie :: \ Minterrupts \ Minterrupts \ Sinterrupts \((register_value),(Minterrupts),(exception))monad \ where \ lift_sie (o1 :: Minterrupts) (d :: Minterrupts) (s :: Sinterrupts) = ( (let (m :: Minterrupts) = o1 in (let m = (if (((((get_Minterrupts_SEI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_SEI m ((get_Sinterrupts_SEI s :: 1 Word.word)) else m) in (let m = (if (((((get_Minterrupts_STI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_STI m ((get_Sinterrupts_STI s :: 1 Word.word)) else m) in (let m = (if (((((get_Minterrupts_SSI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_SSI m ((get_Sinterrupts_SSI s :: 1 Word.word)) else m) in haveNExt () \ ((\ (w__0 :: bool) . return (if w__0 then (let m = (if (((((get_Minterrupts_UEI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_UEI m ((get_Sinterrupts_UEI s :: 1 Word.word)) else m) in (let m = (if (((((get_Minterrupts_UTI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_UTI m ((get_Sinterrupts_UTI s :: 1 Word.word)) else m) in if (((((get_Minterrupts_USI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Minterrupts_USI m ((get_Sinterrupts_USI s :: 1 Word.word)) else m)) else m))))))))\ for o1 :: " Minterrupts " and d :: " Minterrupts " and s :: " Sinterrupts " \ \\val legalize_sie : Minterrupts -> Minterrupts -> mword ty32 -> M Minterrupts\\ definition legalize_sie :: \ Minterrupts \ Minterrupts \(32)Word.word \((register_value),(Minterrupts),(exception))monad \ where \ legalize_sie (m :: Minterrupts) (d :: Minterrupts) (v :: xlenbits) = ( lift_sie m d ((Mk_Sinterrupts v)))\ for m :: " Minterrupts " and d :: " Minterrupts " and v :: "(32)Word.word " \ \\val undefined_Satp64 : unit -> M Satp64\\ definition undefined_Satp64 :: \ unit \((register_value),(Satp64),(exception))monad \ where \ undefined_Satp64 _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . return ((| Satp64_bits = w__0 |)))))\ \ \\val Mk_Satp64 : mword ty64 -> Satp64\\ definition Mk_Satp64 :: \(64)Word.word \ Satp64 \ where \ Mk_Satp64 v = ( (| Satp64_bits = v |) )\ for v :: "(64)Word.word " definition get_Satp64_bits :: \ Satp64 \(64)Word.word \ where \ get_Satp64_bits v = ( (subrange_vec_dec(Satp64_bits v) (( 63 :: int)::ii) (( 0 :: int)::ii) :: 64 Word.word))\ for v :: " Satp64 " definition set_Satp64_bits :: \((regstate),(register_value),(Satp64))register_ref \(64)Word.word \((register_value),(unit),(exception))monad \ where \ set_Satp64_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Satp64_bits := ((update_subrange_vec_dec(Satp64_bits r) (( 63 :: int)::ii) (( 0 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Satp64))register_ref " and v :: "(64)Word.word " definition update_Satp64_bits :: \ Satp64 \(64)Word.word \ Satp64 \ where \ update_Satp64_bits v x = ( ( v (| Satp64_bits := ((update_subrange_vec_dec(Satp64_bits v) (( 63 :: int)::ii) (( 0 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " Satp64 " and x :: "(64)Word.word " \ \\val _get_Satp64_Asid : Satp64 -> mword ty16\\ definition get_Satp64_Asid :: \ Satp64 \(16)Word.word \ where \ get_Satp64_Asid v = ( (subrange_vec_dec(Satp64_bits v) (( 59 :: int)::ii) (( 44 :: int)::ii) :: 16 Word.word))\ for v :: " Satp64 " \ \\val _set_Satp64_Asid : register_ref regstate register_value Satp64 -> mword ty16 -> M unit\\ definition set_Satp64_Asid :: \((regstate),(register_value),(Satp64))register_ref \(16)Word.word \((register_value),(unit),(exception))monad \ where \ set_Satp64_Asid r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Satp64_bits := ((update_subrange_vec_dec(Satp64_bits r) (( 59 :: int)::ii) (( 44 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Satp64))register_ref " and v :: "(16)Word.word " \ \\val _update_Satp64_Asid : Satp64 -> mword ty16 -> Satp64\\ definition update_Satp64_Asid :: \ Satp64 \(16)Word.word \ Satp64 \ where \ update_Satp64_Asid v x = ( ( v (| Satp64_bits := ((update_subrange_vec_dec(Satp64_bits v) (( 59 :: int)::ii) (( 44 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " Satp64 " and x :: "(16)Word.word " \ \\val _update_Satp32_Asid : Satp32 -> mword ty9 -> Satp32\\ \ \\val _get_Satp32_Asid : Satp32 -> mword ty9\\ \ \\val _set_Satp32_Asid : register_ref regstate register_value Satp32 -> mword ty9 -> M unit\\ definition get_Satp64_Mode :: \ Satp64 \(4)Word.word \ where \ get_Satp64_Mode v = ( (subrange_vec_dec(Satp64_bits v) (( 63 :: int)::ii) (( 60 :: int)::ii) :: 4 Word.word))\ for v :: " Satp64 " definition set_Satp64_Mode :: \((regstate),(register_value),(Satp64))register_ref \(4)Word.word \((register_value),(unit),(exception))monad \ where \ set_Satp64_Mode r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Satp64_bits := ((update_subrange_vec_dec(Satp64_bits r) (( 63 :: int)::ii) (( 60 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Satp64))register_ref " and v :: "(4)Word.word " definition update_Satp64_Mode :: \ Satp64 \(4)Word.word \ Satp64 \ where \ update_Satp64_Mode v x = ( ( v (| Satp64_bits := ((update_subrange_vec_dec(Satp64_bits v) (( 63 :: int)::ii) (( 60 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " Satp64 " and x :: "(4)Word.word " \ \\val _get_Satp64_PPN : Satp64 -> mword ty44\\ definition get_Satp64_PPN :: \ Satp64 \(44)Word.word \ where \ get_Satp64_PPN v = ( (subrange_vec_dec(Satp64_bits v) (( 43 :: int)::ii) (( 0 :: int)::ii) :: 44 Word.word))\ for v :: " Satp64 " \ \\val _set_Satp64_PPN : register_ref regstate register_value Satp64 -> mword ty44 -> M unit\\ definition set_Satp64_PPN :: \((regstate),(register_value),(Satp64))register_ref \(44)Word.word \((register_value),(unit),(exception))monad \ where \ set_Satp64_PPN r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Satp64_bits := ((update_subrange_vec_dec(Satp64_bits r) (( 43 :: int)::ii) (( 0 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Satp64))register_ref " and v :: "(44)Word.word " \ \\val _update_Satp64_PPN : Satp64 -> mword ty44 -> Satp64\\ definition update_Satp64_PPN :: \ Satp64 \(44)Word.word \ Satp64 \ where \ update_Satp64_PPN v x = ( ( v (| Satp64_bits := ((update_subrange_vec_dec(Satp64_bits v) (( 43 :: int)::ii) (( 0 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " Satp64 " and x :: "(44)Word.word " \ \\val _update_Satp32_PPN : Satp32 -> mword ty22 -> Satp32\\ \ \\val _get_Satp32_PPN : Satp32 -> mword ty22\\ \ \\val _set_Satp32_PPN : register_ref regstate register_value Satp32 -> mword ty22 -> M unit\\ \ \\val legalize_satp64 : Architecture -> mword ty64 -> mword ty64 -> mword ty64\\ definition legalize_satp64 :: \ Architecture \(64)Word.word \(64)Word.word \(64)Word.word \ where \ legalize_satp64 (a :: Architecture) (o1 :: 64 bits) (v :: 64 bits) = ( (let s = (Mk_Satp64 v) in (case ((satp64Mode_of_bits a ((get_Satp64_Mode s :: 4 Word.word)))) of None => o1 | Some (Sv32) => o1 | Some (_) => (get_Satp64_bits s :: 64 Word.word) )))\ for a :: " Architecture " and o1 :: "(64)Word.word " and v :: "(64)Word.word " \ \\val undefined_Satp32 : unit -> M Satp32\\ definition undefined_Satp32 :: \ unit \((register_value),(Satp32),(exception))monad \ where \ undefined_Satp32 _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Satp32_bits = w__0 |)))))\ \ \\val Mk_Satp32 : mword ty32 -> Satp32\\ definition Mk_Satp32 :: \(32)Word.word \ Satp32 \ where \ Mk_Satp32 v = ( (| Satp32_bits = v |) )\ for v :: "(32)Word.word " definition get_Satp32_bits :: \ Satp32 \(32)Word.word \ where \ get_Satp32_bits v = ( (subrange_vec_dec(Satp32_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Satp32 " definition set_Satp32_bits :: \((regstate),(register_value),(Satp32))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Satp32_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Satp32_bits := ((update_subrange_vec_dec(Satp32_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Satp32))register_ref " and v :: "(32)Word.word " definition update_Satp32_bits :: \ Satp32 \(32)Word.word \ Satp32 \ where \ update_Satp32_bits v x = ( ( v (| Satp32_bits := ((update_subrange_vec_dec(Satp32_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Satp32 " and x :: "(32)Word.word " definition get_Satp32_Asid :: \ Satp32 \(9)Word.word \ where \ get_Satp32_Asid v = ( (subrange_vec_dec(Satp32_bits v) (( 30 :: int)::ii) (( 22 :: int)::ii) :: 9 Word.word))\ for v :: " Satp32 " definition set_Satp32_Asid :: \((regstate),(register_value),(Satp32))register_ref \(9)Word.word \((register_value),(unit),(exception))monad \ where \ set_Satp32_Asid r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Satp32_bits := ((update_subrange_vec_dec(Satp32_bits r) (( 30 :: int)::ii) (( 22 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Satp32))register_ref " and v :: "(9)Word.word " definition update_Satp32_Asid :: \ Satp32 \(9)Word.word \ Satp32 \ where \ update_Satp32_Asid v x = ( ( v (| Satp32_bits := ((update_subrange_vec_dec(Satp32_bits v) (( 30 :: int)::ii) (( 22 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Satp32 " and x :: "(9)Word.word " definition get_Satp32_Mode :: \ Satp32 \(1)Word.word \ where \ get_Satp32_Mode v = ( (subrange_vec_dec(Satp32_bits v) (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word))\ for v :: " Satp32 " definition set_Satp32_Mode :: \((regstate),(register_value),(Satp32))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Satp32_Mode r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Satp32_bits := ((update_subrange_vec_dec(Satp32_bits r) (( 31 :: int)::ii) (( 31 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Satp32))register_ref " and v :: "(1)Word.word " definition update_Satp32_Mode :: \ Satp32 \(1)Word.word \ Satp32 \ where \ update_Satp32_Mode v x = ( ( v (| Satp32_bits := ((update_subrange_vec_dec(Satp32_bits v) (( 31 :: int)::ii) (( 31 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Satp32 " and x :: "(1)Word.word " definition get_Satp32_PPN :: \ Satp32 \(22)Word.word \ where \ get_Satp32_PPN v = ( (subrange_vec_dec(Satp32_bits v) (( 21 :: int)::ii) (( 0 :: int)::ii) :: 22 Word.word))\ for v :: " Satp32 " definition set_Satp32_PPN :: \((regstate),(register_value),(Satp32))register_ref \(22)Word.word \((register_value),(unit),(exception))monad \ where \ set_Satp32_PPN r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Satp32_bits := ((update_subrange_vec_dec(Satp32_bits r) (( 21 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Satp32))register_ref " and v :: "(22)Word.word " definition update_Satp32_PPN :: \ Satp32 \(22)Word.word \ Satp32 \ where \ update_Satp32_PPN v x = ( ( v (| Satp32_bits := ((update_subrange_vec_dec(Satp32_bits v) (( 21 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Satp32 " and x :: "(22)Word.word " \ \\val legalize_satp32 : Architecture -> mword ty32 -> mword ty32 -> mword ty32\\ definition legalize_satp32 :: \ Architecture \(32)Word.word \(32)Word.word \(32)Word.word \ where \ legalize_satp32 (a :: Architecture) (o1 :: 32 bits) (v :: 32 bits) = ( v )\ for a :: " Architecture " and o1 :: "(32)Word.word " and v :: "(32)Word.word " \ \\val PmpAddrMatchType_of_num : integer -> PmpAddrMatchType\\ definition PmpAddrMatchType_of_num :: \ int \ PmpAddrMatchType \ where \ PmpAddrMatchType_of_num arg1 = ( (let l__89 = arg1 in if (((l__89 = (( 0 :: int)::ii)))) then OFF else if (((l__89 = (( 1 :: int)::ii)))) then TOR else if (((l__89 = (( 2 :: int)::ii)))) then NA4 else NAPOT))\ for arg1 :: " int " \ \\val num_of_PmpAddrMatchType : PmpAddrMatchType -> integer\\ fun num_of_PmpAddrMatchType :: \ PmpAddrMatchType \ int \ where \ num_of_PmpAddrMatchType OFF = ( (( 0 :: int)::ii))\ |\ num_of_PmpAddrMatchType TOR = ( (( 1 :: int)::ii))\ |\ num_of_PmpAddrMatchType NA4 = ( (( 2 :: int)::ii))\ |\ num_of_PmpAddrMatchType NAPOT = ( (( 3 :: int)::ii))\ \ \\val pmpAddrMatchType_of_bits : mword ty2 -> M PmpAddrMatchType\\ definition pmpAddrMatchType_of_bits :: \(2)Word.word \((register_value),(PmpAddrMatchType),(exception))monad \ where \ pmpAddrMatchType_of_bits bs = ( (let b__0 = bs in if (((b__0 = ( 0b00 :: 2 Word.word)))) then return OFF else if (((b__0 = ( 0b01 :: 2 Word.word)))) then return TOR else if (((b__0 = ( 0b10 :: 2 Word.word)))) then return NA4 else if (((b__0 = ( 0b11 :: 2 Word.word)))) then return NAPOT else assert_exp False (''Pattern match failure at model/riscv_pmp_regs.sail 7:2 - 12:3'') \ exit0 () ))\ for bs :: "(2)Word.word " \ \\val pmpAddrMatchType_to_bits : PmpAddrMatchType -> mword ty2\\ fun pmpAddrMatchType_to_bits :: \ PmpAddrMatchType \(2)Word.word \ where \ pmpAddrMatchType_to_bits OFF = ( ( 0b00 :: 2 Word.word))\ |\ pmpAddrMatchType_to_bits TOR = ( ( 0b01 :: 2 Word.word))\ |\ pmpAddrMatchType_to_bits NA4 = ( ( 0b10 :: 2 Word.word))\ |\ pmpAddrMatchType_to_bits NAPOT = ( ( 0b11 :: 2 Word.word))\ \ \\val undefined_Pmpcfg_ent : unit -> M Pmpcfg_ent\\ definition undefined_Pmpcfg_ent :: \ unit \((register_value),(Pmpcfg_ent),(exception))monad \ where \ undefined_Pmpcfg_ent _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 8 Word.word) M) \ ((\ (w__0 :: 8 Word.word) . return ((| Pmpcfg_ent_bits = w__0 |)))))\ \ \\val Mk_Pmpcfg_ent : mword ty8 -> Pmpcfg_ent\\ definition Mk_Pmpcfg_ent :: \(8)Word.word \ Pmpcfg_ent \ where \ Mk_Pmpcfg_ent v = ( (| Pmpcfg_ent_bits = v |) )\ for v :: "(8)Word.word " definition get_Pmpcfg_ent_bits :: \ Pmpcfg_ent \(8)Word.word \ where \ get_Pmpcfg_ent_bits v = ( (subrange_vec_dec(Pmpcfg_ent_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word))\ for v :: " Pmpcfg_ent " definition set_Pmpcfg_ent_bits :: \((regstate),(register_value),(Pmpcfg_ent))register_ref \(8)Word.word \((register_value),(unit),(exception))monad \ where \ set_Pmpcfg_ent_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits r) (( 7 :: int)::ii) (( 0 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Pmpcfg_ent))register_ref " and v :: "(8)Word.word " definition update_Pmpcfg_ent_bits :: \ Pmpcfg_ent \(8)Word.word \ Pmpcfg_ent \ where \ update_Pmpcfg_ent_bits v x = ( ( v (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " Pmpcfg_ent " and x :: "(8)Word.word " definition get_Pmpcfg_ent_A :: \ Pmpcfg_ent \(2)Word.word \ where \ get_Pmpcfg_ent_A v = ( (subrange_vec_dec(Pmpcfg_ent_bits v) (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word))\ for v :: " Pmpcfg_ent " definition set_Pmpcfg_ent_A :: \((regstate),(register_value),(Pmpcfg_ent))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_Pmpcfg_ent_A r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits r) (( 4 :: int)::ii) (( 3 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Pmpcfg_ent))register_ref " and v :: "(2)Word.word " definition update_Pmpcfg_ent_A :: \ Pmpcfg_ent \(2)Word.word \ Pmpcfg_ent \ where \ update_Pmpcfg_ent_A v x = ( ( v (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits v) (( 4 :: int)::ii) (( 3 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " Pmpcfg_ent " and x :: "(2)Word.word " definition get_Pmpcfg_ent_L :: \ Pmpcfg_ent \(1)Word.word \ where \ get_Pmpcfg_ent_L v = ( (subrange_vec_dec(Pmpcfg_ent_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word))\ for v :: " Pmpcfg_ent " definition set_Pmpcfg_ent_L :: \((regstate),(register_value),(Pmpcfg_ent))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Pmpcfg_ent_L r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits r) (( 7 :: int)::ii) (( 7 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Pmpcfg_ent))register_ref " and v :: "(1)Word.word " definition update_Pmpcfg_ent_L :: \ Pmpcfg_ent \(1)Word.word \ Pmpcfg_ent \ where \ update_Pmpcfg_ent_L v x = ( ( v (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " Pmpcfg_ent " and x :: "(1)Word.word " definition get_Pmpcfg_ent_R :: \ Pmpcfg_ent \(1)Word.word \ where \ get_Pmpcfg_ent_R v = ( (subrange_vec_dec(Pmpcfg_ent_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Pmpcfg_ent " definition set_Pmpcfg_ent_R :: \((regstate),(register_value),(Pmpcfg_ent))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Pmpcfg_ent_R r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Pmpcfg_ent))register_ref " and v :: "(1)Word.word " definition update_Pmpcfg_ent_R :: \ Pmpcfg_ent \(1)Word.word \ Pmpcfg_ent \ where \ update_Pmpcfg_ent_R v x = ( ( v (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " Pmpcfg_ent " and x :: "(1)Word.word " definition get_Pmpcfg_ent_W :: \ Pmpcfg_ent \(1)Word.word \ where \ get_Pmpcfg_ent_W v = ( (subrange_vec_dec(Pmpcfg_ent_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " Pmpcfg_ent " definition set_Pmpcfg_ent_W :: \((regstate),(register_value),(Pmpcfg_ent))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Pmpcfg_ent_W r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Pmpcfg_ent))register_ref " and v :: "(1)Word.word " definition update_Pmpcfg_ent_W :: \ Pmpcfg_ent \(1)Word.word \ Pmpcfg_ent \ where \ update_Pmpcfg_ent_W v x = ( ( v (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " Pmpcfg_ent " and x :: "(1)Word.word " definition get_Pmpcfg_ent_X :: \ Pmpcfg_ent \(1)Word.word \ where \ get_Pmpcfg_ent_X v = ( (subrange_vec_dec(Pmpcfg_ent_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word))\ for v :: " Pmpcfg_ent " definition set_Pmpcfg_ent_X :: \((regstate),(register_value),(Pmpcfg_ent))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Pmpcfg_ent_X r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits r) (( 2 :: int)::ii) (( 2 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Pmpcfg_ent))register_ref " and v :: "(1)Word.word " definition update_Pmpcfg_ent_X :: \ Pmpcfg_ent \(1)Word.word \ Pmpcfg_ent \ where \ update_Pmpcfg_ent_X v x = ( ( v (| Pmpcfg_ent_bits := ((update_subrange_vec_dec(Pmpcfg_ent_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " Pmpcfg_ent " and x :: "(1)Word.word " \ \\val pmpReadCfgReg : integer -> M (mword ty32)\\ definition pmpReadCfgReg :: \ int \((register_value),((32)Word.word),(exception))monad \ where \ pmpReadCfgReg n = ( (let l__85 = n in if (((l__85 = (( 0 :: int)::ii)))) then read_reg pmp3cfg_ref \ ((\ (w__0 :: Pmpcfg_ent) . read_reg pmp2cfg_ref \ ((\ (w__1 :: Pmpcfg_ent) . read_reg pmp1cfg_ref \ ((\ (w__2 :: Pmpcfg_ent) . read_reg pmp0cfg_ref \ ((\ (w__3 :: Pmpcfg_ent) . return ((concat_vec ((get_Pmpcfg_ent_bits w__0 :: 8 Word.word)) ((concat_vec ((get_Pmpcfg_ent_bits w__1 :: 8 Word.word)) ((concat_vec ((get_Pmpcfg_ent_bits w__2 :: 8 Word.word)) ((get_Pmpcfg_ent_bits w__3 :: 8 Word.word)) :: 16 Word.word)) :: 24 Word.word)) :: 32 Word.word)))))))))) else if (((l__85 = (( 1 :: int)::ii)))) then read_reg pmp7cfg_ref \ ((\ (w__4 :: Pmpcfg_ent) . read_reg pmp6cfg_ref \ ((\ (w__5 :: Pmpcfg_ent) . read_reg pmp5cfg_ref \ ((\ (w__6 :: Pmpcfg_ent) . read_reg pmp4cfg_ref \ ((\ (w__7 :: Pmpcfg_ent) . return ((concat_vec ((get_Pmpcfg_ent_bits w__4 :: 8 Word.word)) ((concat_vec ((get_Pmpcfg_ent_bits w__5 :: 8 Word.word)) ((concat_vec ((get_Pmpcfg_ent_bits w__6 :: 8 Word.word)) ((get_Pmpcfg_ent_bits w__7 :: 8 Word.word)) :: 16 Word.word)) :: 24 Word.word)) :: 32 Word.word)))))))))) else if (((l__85 = (( 2 :: int)::ii)))) then read_reg pmp11cfg_ref \ ((\ (w__8 :: Pmpcfg_ent) . read_reg pmp10cfg_ref \ ((\ (w__9 :: Pmpcfg_ent) . read_reg pmp9cfg_ref \ ((\ (w__10 :: Pmpcfg_ent) . read_reg pmp8cfg_ref \ ((\ (w__11 :: Pmpcfg_ent) . return ((concat_vec ((get_Pmpcfg_ent_bits w__8 :: 8 Word.word)) ((concat_vec ((get_Pmpcfg_ent_bits w__9 :: 8 Word.word)) ((concat_vec ((get_Pmpcfg_ent_bits w__10 :: 8 Word.word)) ((get_Pmpcfg_ent_bits w__11 :: 8 Word.word)) :: 16 Word.word)) :: 24 Word.word)) :: 32 Word.word)))))))))) else if (((l__85 = (( 3 :: int)::ii)))) then read_reg pmp15cfg_ref \ ((\ (w__12 :: Pmpcfg_ent) . read_reg pmp14cfg_ref \ ((\ (w__13 :: Pmpcfg_ent) . read_reg pmp13cfg_ref \ ((\ (w__14 :: Pmpcfg_ent) . read_reg pmp12cfg_ref \ ((\ (w__15 :: Pmpcfg_ent) . return ((concat_vec ((get_Pmpcfg_ent_bits w__12 :: 8 Word.word)) ((concat_vec ((get_Pmpcfg_ent_bits w__13 :: 8 Word.word)) ((concat_vec ((get_Pmpcfg_ent_bits w__14 :: 8 Word.word)) ((get_Pmpcfg_ent_bits w__15 :: 8 Word.word)) :: 16 Word.word)) :: 24 Word.word)) :: 32 Word.word)))))))))) else assert_exp False (''Pattern match failure at model/riscv_pmp_regs.sail 75:2 - 85:8'') \ exit0 () ))\ for n :: " int " \ \\val pmpLocked : Pmpcfg_ent -> bool\\ definition pmpLocked :: \ Pmpcfg_ent \ bool \ where \ pmpLocked cfg = ( (((get_Pmpcfg_ent_L cfg :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))\ for cfg :: " Pmpcfg_ent " \ \\val pmpTORLocked : Pmpcfg_ent -> M bool\\ definition pmpTORLocked :: \ Pmpcfg_ent \((register_value),(bool),(exception))monad \ where \ pmpTORLocked cfg = ( and_boolM (return (((((get_Pmpcfg_ent_L cfg :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))) (pmpAddrMatchType_of_bits ((get_Pmpcfg_ent_A cfg :: 2 Word.word)) \ ((\ (w__0 :: PmpAddrMatchType) . return (((w__0 = TOR)))))))\ for cfg :: " Pmpcfg_ent " \ \\val pmpWriteCfg : Pmpcfg_ent -> mword ty8 -> Pmpcfg_ent\\ definition pmpWriteCfg :: \ Pmpcfg_ent \(8)Word.word \ Pmpcfg_ent \ where \ pmpWriteCfg (cfg :: Pmpcfg_ent) (v :: 8 bits) = ( if ((pmpLocked cfg)) then cfg else Mk_Pmpcfg_ent ((and_vec v ( 0x9F :: 8 Word.word) :: 8 Word.word)))\ for cfg :: " Pmpcfg_ent " and v :: "(8)Word.word " \ \\val pmpWriteCfgReg : integer -> mword ty32 -> M unit\\ definition pmpWriteCfgReg :: \ int \(32)Word.word \((register_value),(unit),(exception))monad \ where \ pmpWriteCfgReg n v = ( (let l__81 = n in if (((l__81 = (( 0 :: int)::ii)))) then read_reg pmp0cfg_ref \ ((\ (w__0 :: Pmpcfg_ent) . (write_reg pmp0cfg_ref ((pmpWriteCfg w__0 ((subrange_vec_dec v (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp1cfg_ref) \ ((\ (w__1 :: Pmpcfg_ent) . (write_reg pmp1cfg_ref ((pmpWriteCfg w__1 ((subrange_vec_dec v (( 15 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp2cfg_ref) \ ((\ (w__2 :: Pmpcfg_ent) . (write_reg pmp2cfg_ref ((pmpWriteCfg w__2 ((subrange_vec_dec v (( 23 :: int)::ii) (( 16 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp3cfg_ref) \ ((\ (w__3 :: Pmpcfg_ent) . write_reg pmp3cfg_ref ((pmpWriteCfg w__3 ((subrange_vec_dec v (( 31 :: int)::ii) (( 24 :: int)::ii) :: 8 Word.word)))))))))))) else if (((l__81 = (( 1 :: int)::ii)))) then read_reg pmp4cfg_ref \ ((\ (w__4 :: Pmpcfg_ent) . (write_reg pmp4cfg_ref ((pmpWriteCfg w__4 ((subrange_vec_dec v (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp5cfg_ref) \ ((\ (w__5 :: Pmpcfg_ent) . (write_reg pmp5cfg_ref ((pmpWriteCfg w__5 ((subrange_vec_dec v (( 15 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp6cfg_ref) \ ((\ (w__6 :: Pmpcfg_ent) . (write_reg pmp6cfg_ref ((pmpWriteCfg w__6 ((subrange_vec_dec v (( 23 :: int)::ii) (( 16 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp7cfg_ref) \ ((\ (w__7 :: Pmpcfg_ent) . write_reg pmp7cfg_ref ((pmpWriteCfg w__7 ((subrange_vec_dec v (( 31 :: int)::ii) (( 24 :: int)::ii) :: 8 Word.word)))))))))))) else if (((l__81 = (( 2 :: int)::ii)))) then read_reg pmp8cfg_ref \ ((\ (w__8 :: Pmpcfg_ent) . (write_reg pmp8cfg_ref ((pmpWriteCfg w__8 ((subrange_vec_dec v (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp9cfg_ref) \ ((\ (w__9 :: Pmpcfg_ent) . (write_reg pmp9cfg_ref ((pmpWriteCfg w__9 ((subrange_vec_dec v (( 15 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp10cfg_ref) \ ((\ (w__10 :: Pmpcfg_ent) . (write_reg pmp10cfg_ref ((pmpWriteCfg w__10 ((subrange_vec_dec v (( 23 :: int)::ii) (( 16 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp11cfg_ref) \ ((\ (w__11 :: Pmpcfg_ent) . write_reg pmp11cfg_ref ((pmpWriteCfg w__11 ((subrange_vec_dec v (( 31 :: int)::ii) (( 24 :: int)::ii) :: 8 Word.word)))))))))))) else if (((l__81 = (( 3 :: int)::ii)))) then read_reg pmp12cfg_ref \ ((\ (w__12 :: Pmpcfg_ent) . (write_reg pmp12cfg_ref ((pmpWriteCfg w__12 ((subrange_vec_dec v (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp13cfg_ref) \ ((\ (w__13 :: Pmpcfg_ent) . (write_reg pmp13cfg_ref ((pmpWriteCfg w__13 ((subrange_vec_dec v (( 15 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp14cfg_ref) \ ((\ (w__14 :: Pmpcfg_ent) . (write_reg pmp14cfg_ref ((pmpWriteCfg w__14 ((subrange_vec_dec v (( 23 :: int)::ii) (( 16 :: int)::ii) :: 8 Word.word)))) \ read_reg pmp15cfg_ref) \ ((\ (w__15 :: Pmpcfg_ent) . write_reg pmp15cfg_ref ((pmpWriteCfg w__15 ((subrange_vec_dec v (( 31 :: int)::ii) (( 24 :: int)::ii) :: 8 Word.word)))))))))))) else assert_exp False (''Pattern match failure at model/riscv_pmp_regs.sail 101:2 - 144:8'') \ exit0 () ))\ for n :: " int " and v :: "(32)Word.word " \ \\val pmpWriteAddr : bool -> bool -> mword ty32 -> mword ty32 -> mword ty32\\ definition pmpWriteAddr :: \ bool \ bool \(32)Word.word \(32)Word.word \(32)Word.word \ where \ pmpWriteAddr (locked :: bool) (tor_locked :: bool) (reg :: xlenbits) (v :: xlenbits) = ( if (((locked \ tor_locked))) then reg else v )\ for locked :: " bool " and tor_locked :: " bool " and reg :: "(32)Word.word " and v :: "(32)Word.word " \ \\val pmpAddrRange : Pmpcfg_ent -> mword ty32 -> mword ty32 -> M (maybe ((mword ty32 * mword ty32)))\\ definition pmpAddrRange :: \ Pmpcfg_ent \(32)Word.word \(32)Word.word \((register_value),(((32)Word.word*(32)Word.word)option),(exception))monad \ where \ pmpAddrRange (cfg :: Pmpcfg_ent) (pmpaddr :: xlenbits) (prev_pmpaddr :: xlenbits) = ( pmpAddrMatchType_of_bits ((get_Pmpcfg_ent_A cfg :: 2 Word.word)) \ ((\ (w__0 :: PmpAddrMatchType) . return ((case w__0 of OFF => None | TOR => Some ((shiftl prev_pmpaddr (( 2 :: int)::ii) :: 32 Word.word), (shiftl pmpaddr (( 2 :: int)::ii) :: 32 Word.word)) | NA4 => (let lo = ((shiftl pmpaddr (( 2 :: int)::ii) :: 32 Word.word)) in Some (lo, (add_vec_int lo (( 4 :: int)::ii) :: 32 Word.word))) | NAPOT => (let mask1 = ((xor_vec pmpaddr ((add_vec_int pmpaddr (( 1 :: int)::ii) :: 32 Word.word)) :: 32 Word.word)) in (let lo = ((and_vec pmpaddr ((not_vec mask1 :: 32 Word.word)) :: 32 Word.word)) in (let len = ((add_vec_int mask1 (( 1 :: int)::ii) :: 32 Word.word)) in Some ((shiftl lo (( 2 :: int)::ii) :: 32 Word.word), (shiftl ((add_vec lo len :: 32 Word.word)) (( 2 :: int)::ii) :: 32 Word.word))))) )))))\ for cfg :: " Pmpcfg_ent " and pmpaddr :: "(32)Word.word " and prev_pmpaddr :: "(32)Word.word " \ \\val pmpCheckRWX : Pmpcfg_ent -> AccessType unit -> bool\\ fun pmpCheckRWX :: \ Pmpcfg_ent \(unit)AccessType \ bool \ where \ pmpCheckRWX ent (Read (_)) = ( (((get_Pmpcfg_ent_R ent :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))\ for ent :: " Pmpcfg_ent " |\ pmpCheckRWX ent (Write (_)) = ( (((get_Pmpcfg_ent_W ent :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))\ for ent :: " Pmpcfg_ent " |\ pmpCheckRWX ent (ReadWrite (_)) = ( ((((((get_Pmpcfg_ent_R ent :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ (((((get_Pmpcfg_ent_W ent :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))\ for ent :: " Pmpcfg_ent " |\ pmpCheckRWX ent (Execute (_)) = ( (((get_Pmpcfg_ent_X ent :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))\ for ent :: " Pmpcfg_ent " \ \\val pmpCheckPerms : Pmpcfg_ent -> AccessType unit -> Privilege -> bool\\ fun pmpCheckPerms :: \ Pmpcfg_ent \(unit)AccessType \ Privilege \ bool \ where \ pmpCheckPerms ent acc1 Machine = ( if ((pmpLocked ent)) then pmpCheckRWX ent acc1 else True )\ for ent :: " Pmpcfg_ent " and acc1 :: "(unit)AccessType " |\ pmpCheckPerms ent acc1 _ = ( pmpCheckRWX ent acc1 )\ for ent :: " Pmpcfg_ent " and acc1 :: "(unit)AccessType " \ \\val pmpAddrMatch_of_num : integer -> pmpAddrMatch\\ definition pmpAddrMatch_of_num :: \ int \ pmpAddrMatch \ where \ pmpAddrMatch_of_num arg1 = ( (let l__79 = arg1 in if (((l__79 = (( 0 :: int)::ii)))) then PMP_NoMatch else if (((l__79 = (( 1 :: int)::ii)))) then PMP_PartialMatch else PMP_Match))\ for arg1 :: " int " \ \\val num_of_pmpAddrMatch : pmpAddrMatch -> integer\\ fun num_of_pmpAddrMatch :: \ pmpAddrMatch \ int \ where \ num_of_pmpAddrMatch PMP_NoMatch = ( (( 0 :: int)::ii))\ |\ num_of_pmpAddrMatch PMP_PartialMatch = ( (( 1 :: int)::ii))\ |\ num_of_pmpAddrMatch PMP_Match = ( (( 2 :: int)::ii))\ \ \\val pmpMatchAddr : mword ty32 -> mword ty32 -> maybe ((mword ty32 * mword ty32)) -> pmpAddrMatch\\ fun pmpMatchAddr :: \(32)Word.word \(32)Word.word \(xlenbits*xlenbits)option \ pmpAddrMatch \ where \ pmpMatchAddr (addr :: xlenbits) (width :: xlenbits) (None :: pmp_addr_range) = ( PMP_NoMatch )\ for addr :: "(32)Word.word " and width :: "(32)Word.word " |\ pmpMatchAddr (addr :: xlenbits) (width :: xlenbits) ((Some ((lo, hi))) :: pmp_addr_range) = ( if ((zopz0zI_u hi lo)) then PMP_NoMatch else if (((((zopz0zIzJ_u ((add_vec addr width :: 32 Word.word)) lo)) \ ((zopz0zIzJ_u hi addr))))) then PMP_NoMatch else if (((((zopz0zIzJ_u lo addr)) \ ((zopz0zIzJ_u ((add_vec addr width :: 32 Word.word)) hi))))) then PMP_Match else PMP_PartialMatch )\ for addr :: "(32)Word.word " and width :: "(32)Word.word " and hi :: "(32)Word.word " and lo :: "(32)Word.word " \ \\val pmpMatch_of_num : integer -> pmpMatch\\ definition pmpMatch_of_num :: \ int \ pmpMatch \ where \ pmpMatch_of_num arg1 = ( (let l__77 = arg1 in if (((l__77 = (( 0 :: int)::ii)))) then PMP_Success else if (((l__77 = (( 1 :: int)::ii)))) then PMP_Continue else PMP_Fail))\ for arg1 :: " int " \ \\val num_of_pmpMatch : pmpMatch -> integer\\ fun num_of_pmpMatch :: \ pmpMatch \ int \ where \ num_of_pmpMatch PMP_Success = ( (( 0 :: int)::ii))\ |\ num_of_pmpMatch PMP_Continue = ( (( 1 :: int)::ii))\ |\ num_of_pmpMatch PMP_Fail = ( (( 2 :: int)::ii))\ \ \\val pmpMatchEntry : mword ty32 -> mword ty32 -> AccessType unit -> Privilege -> Pmpcfg_ent -> mword ty32 -> mword ty32 -> M pmpMatch\\ definition pmpMatchEntry :: \(32)Word.word \(32)Word.word \(ext_access_type)AccessType \ Privilege \ Pmpcfg_ent \(32)Word.word \(32)Word.word \((register_value),(pmpMatch),(exception))monad \ where \ pmpMatchEntry (addr :: xlenbits) (width :: xlenbits) (acc1 :: ext_access_type AccessType) (priv :: Privilege) (ent :: Pmpcfg_ent) (pmpaddr :: xlenbits) (prev_pmpaddr :: xlenbits) = ( (pmpAddrRange ent pmpaddr prev_pmpaddr :: ( (( 32 Word.word * 32 Word.word))option) M) \ ((\ rng . return ((case ((pmpMatchAddr addr width rng)) of PMP_NoMatch => PMP_Continue | PMP_PartialMatch => PMP_Fail | PMP_Match => if ((pmpCheckPerms ent acc1 priv)) then PMP_Success else PMP_Fail )))))\ for addr :: "(32)Word.word " and width :: "(32)Word.word " and acc1 :: "(ext_access_type)AccessType " and priv :: " Privilege " and ent :: " Pmpcfg_ent " and pmpaddr :: "(32)Word.word " and prev_pmpaddr :: "(32)Word.word " \ \\val pmpCheck : mword ty32 -> integer -> AccessType unit -> Privilege -> M (maybe ExceptionType)\\ definition pmpCheck :: \(32)Word.word \ int \(ext_access_type)AccessType \ Privilege \((register_value),((ExceptionType)option),(exception))monad \ where \ pmpCheck (addr :: xlenbits) (width :: int) (acc1 :: ext_access_type AccessType) (priv :: Privilege) = ( (let (width :: xlenbits) = ((to_bits (( 32 :: int)::ii) width :: 32 Word.word)) in read_reg pmp0cfg_ref \ ((\ (w__0 :: Pmpcfg_ent) . (read_reg pmpaddr0_ref :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__0 w__1 ((zeros_implicit (( 32 :: int)::ii) :: 32 Word.word)) \ ((\ (w__2 :: pmpMatch) . (case w__2 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp1cfg_ref \ ((\ (w__3 :: Pmpcfg_ent) . (read_reg pmpaddr1_ref :: ( 32 Word.word) M) \ ((\ (w__4 :: 32 Word.word) . (read_reg pmpaddr0_ref :: ( 32 Word.word) M) \ ((\ (w__5 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__3 w__4 w__5 \ ((\ (w__6 :: pmpMatch) . (case w__6 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp2cfg_ref \ ((\ (w__7 :: Pmpcfg_ent) . (read_reg pmpaddr2_ref :: ( 32 Word.word) M) \ ((\ (w__8 :: 32 Word.word) . (read_reg pmpaddr1_ref :: ( 32 Word.word) M) \ ((\ (w__9 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__7 w__8 w__9 \ ((\ (w__10 :: pmpMatch) . (case w__10 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp3cfg_ref \ ((\ (w__11 :: Pmpcfg_ent) . (read_reg pmpaddr3_ref :: ( 32 Word.word) M) \ ((\ (w__12 :: 32 Word.word) . (read_reg pmpaddr2_ref :: ( 32 Word.word) M) \ ((\ (w__13 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__11 w__12 w__13 \ ((\ (w__14 :: pmpMatch) . (case w__14 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp4cfg_ref \ ((\ (w__15 :: Pmpcfg_ent) . (read_reg pmpaddr4_ref :: ( 32 Word.word) M) \ ((\ (w__16 :: 32 Word.word) . (read_reg pmpaddr3_ref :: ( 32 Word.word) M) \ ((\ (w__17 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__15 w__16 w__17 \ ((\ (w__18 :: pmpMatch) . (case w__18 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp5cfg_ref \ ((\ (w__19 :: Pmpcfg_ent) . (read_reg pmpaddr5_ref :: ( 32 Word.word) M) \ ((\ (w__20 :: 32 Word.word) . (read_reg pmpaddr4_ref :: ( 32 Word.word) M) \ ((\ (w__21 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__19 w__20 w__21 \ ((\ (w__22 :: pmpMatch) . (case w__22 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp6cfg_ref \ ((\ (w__23 :: Pmpcfg_ent) . (read_reg pmpaddr6_ref :: ( 32 Word.word) M) \ ((\ (w__24 :: 32 Word.word) . (read_reg pmpaddr5_ref :: ( 32 Word.word) M) \ ((\ (w__25 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__23 w__24 w__25 \ ((\ (w__26 :: pmpMatch) . (case w__26 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp7cfg_ref \ ((\ (w__27 :: Pmpcfg_ent) . (read_reg pmpaddr7_ref :: ( 32 Word.word) M) \ ((\ (w__28 :: 32 Word.word) . (read_reg pmpaddr6_ref :: ( 32 Word.word) M) \ ((\ (w__29 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__27 w__28 w__29 \ ((\ (w__30 :: pmpMatch) . (case w__30 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp8cfg_ref \ ((\ (w__31 :: Pmpcfg_ent) . (read_reg pmpaddr8_ref :: ( 32 Word.word) M) \ ((\ (w__32 :: 32 Word.word) . (read_reg pmpaddr7_ref :: ( 32 Word.word) M) \ ((\ (w__33 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__31 w__32 w__33 \ ((\ (w__34 :: pmpMatch) . (case w__34 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp9cfg_ref \ ((\ (w__35 :: Pmpcfg_ent) . (read_reg pmpaddr9_ref :: ( 32 Word.word) M) \ ((\ (w__36 :: 32 Word.word) . (read_reg pmpaddr8_ref :: ( 32 Word.word) M) \ ((\ (w__37 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__35 w__36 w__37 \ ((\ (w__38 :: pmpMatch) . (case w__38 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp10cfg_ref \ ((\ (w__39 :: Pmpcfg_ent) . (read_reg pmpaddr10_ref :: ( 32 Word.word) M) \ ((\ (w__40 :: 32 Word.word) . (read_reg pmpaddr9_ref :: ( 32 Word.word) M) \ ((\ (w__41 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__39 w__40 w__41 \ ((\ (w__42 :: pmpMatch) . (case w__42 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp11cfg_ref \ ((\ (w__43 :: Pmpcfg_ent) . (read_reg pmpaddr11_ref :: ( 32 Word.word) M) \ ((\ (w__44 :: 32 Word.word) . (read_reg pmpaddr10_ref :: ( 32 Word.word) M) \ ((\ (w__45 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__43 w__44 w__45 \ ((\ (w__46 :: pmpMatch) . (case w__46 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp12cfg_ref \ ((\ (w__47 :: Pmpcfg_ent) . (read_reg pmpaddr12_ref :: ( 32 Word.word) M) \ ((\ (w__48 :: 32 Word.word) . (read_reg pmpaddr11_ref :: ( 32 Word.word) M) \ ((\ (w__49 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__47 w__48 w__49 \ ((\ (w__50 :: pmpMatch) . (case w__50 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp13cfg_ref \ ((\ (w__51 :: Pmpcfg_ent) . (read_reg pmpaddr13_ref :: ( 32 Word.word) M) \ ((\ (w__52 :: 32 Word.word) . (read_reg pmpaddr12_ref :: ( 32 Word.word) M) \ ((\ (w__53 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__51 w__52 w__53 \ ((\ (w__54 :: pmpMatch) . (case w__54 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp14cfg_ref \ ((\ (w__55 :: Pmpcfg_ent) . (read_reg pmpaddr14_ref :: ( 32 Word.word) M) \ ((\ (w__56 :: 32 Word.word) . (read_reg pmpaddr13_ref :: ( 32 Word.word) M) \ ((\ (w__57 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__55 w__56 w__57 \ ((\ (w__58 :: pmpMatch) . (case w__58 of PMP_Success => return True | PMP_Fail => return False | PMP_Continue => read_reg pmp15cfg_ref \ ((\ (w__59 :: Pmpcfg_ent) . (read_reg pmpaddr15_ref :: ( 32 Word.word) M) \ ((\ (w__60 :: 32 Word.word) . (read_reg pmpaddr14_ref :: ( 32 Word.word) M) \ ((\ (w__61 :: 32 Word.word) . pmpMatchEntry addr width acc1 priv w__59 w__60 w__61 \ ((\ (w__62 :: pmpMatch) . return ((case w__62 of PMP_Success => True | PMP_Fail => False | PMP_Continue => (case priv of Machine => True | _ => False ) )))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ))))))))) ) \ ((\ (check' :: bool) . return (if check' then None else (case acc1 of Read (_) => Some (E_Load_Access_Fault () ) | Write (_) => Some (E_SAMO_Access_Fault () ) | ReadWrite (_) => Some (E_SAMO_Access_Fault () ) | Execute (_) => Some (E_Fetch_Access_Fault () ) ))))))))))))\ for addr :: "(32)Word.word " and width :: " int " and acc1 :: "(ext_access_type)AccessType " and priv :: " Privilege " \ \\val init_pmp : unit -> M unit\\ definition init_pmp :: \ unit \((register_value),(unit),(exception))monad \ where \ init_pmp _ = ( read_reg pmp0cfg_ref \ ((\ (w__0 :: Pmpcfg_ent) . (write_reg pmp0cfg_ref ((update_Pmpcfg_ent_A w__0 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp1cfg_ref) \ ((\ (w__1 :: Pmpcfg_ent) . (write_reg pmp1cfg_ref ((update_Pmpcfg_ent_A w__1 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp2cfg_ref) \ ((\ (w__2 :: Pmpcfg_ent) . (write_reg pmp2cfg_ref ((update_Pmpcfg_ent_A w__2 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp3cfg_ref) \ ((\ (w__3 :: Pmpcfg_ent) . (write_reg pmp3cfg_ref ((update_Pmpcfg_ent_A w__3 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp4cfg_ref) \ ((\ (w__4 :: Pmpcfg_ent) . (write_reg pmp4cfg_ref ((update_Pmpcfg_ent_A w__4 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp5cfg_ref) \ ((\ (w__5 :: Pmpcfg_ent) . (write_reg pmp5cfg_ref ((update_Pmpcfg_ent_A w__5 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp6cfg_ref) \ ((\ (w__6 :: Pmpcfg_ent) . (write_reg pmp6cfg_ref ((update_Pmpcfg_ent_A w__6 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp7cfg_ref) \ ((\ (w__7 :: Pmpcfg_ent) . (write_reg pmp7cfg_ref ((update_Pmpcfg_ent_A w__7 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp8cfg_ref) \ ((\ (w__8 :: Pmpcfg_ent) . (write_reg pmp8cfg_ref ((update_Pmpcfg_ent_A w__8 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp9cfg_ref) \ ((\ (w__9 :: Pmpcfg_ent) . (write_reg pmp9cfg_ref ((update_Pmpcfg_ent_A w__9 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp10cfg_ref) \ ((\ (w__10 :: Pmpcfg_ent) . (write_reg pmp10cfg_ref ((update_Pmpcfg_ent_A w__10 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp11cfg_ref) \ ((\ (w__11 :: Pmpcfg_ent) . (write_reg pmp11cfg_ref ((update_Pmpcfg_ent_A w__11 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp12cfg_ref) \ ((\ (w__12 :: Pmpcfg_ent) . (write_reg pmp12cfg_ref ((update_Pmpcfg_ent_A w__12 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp13cfg_ref) \ ((\ (w__13 :: Pmpcfg_ent) . (write_reg pmp13cfg_ref ((update_Pmpcfg_ent_A w__13 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp14cfg_ref) \ ((\ (w__14 :: Pmpcfg_ent) . (write_reg pmp14cfg_ref ((update_Pmpcfg_ent_A w__14 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))) \ read_reg pmp15cfg_ref) \ ((\ (w__15 :: Pmpcfg_ent) . write_reg pmp15cfg_ref ((update_Pmpcfg_ent_A w__15 ((pmpAddrMatchType_to_bits OFF :: 2 Word.word)))))))))))))))))))))))))))))))))))))\ \ \\val ext_init_regs : unit -> M unit\\ definition ext_init_regs :: \ unit \((register_value),(unit),(exception))monad \ where \ ext_init_regs _ = ( return () )\ \ \\ This function is called after above when running rvfi and allows the model to be initialised differently (e.g. CHERI cap regs are initialised to omnipotent instead of null). \\ \ \\val ext_rvfi_init : unit -> M unit\\ definition ext_rvfi_init :: \ unit \((register_value),(unit),(exception))monad \ where \ ext_rvfi_init _ = ( (read_reg x1_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . write_reg x1_ref w__0)))\ \ \\ THIS(csrno, priv, isWrite) allows an extension to block access to csrno, at Privilege level priv. It should return true if the access is allowed. \\ \ \\val ext_check_CSR : mword ty12 -> Privilege -> bool -> bool\\ definition ext_check_CSR :: \(12)Word.word \ Privilege \ bool \ bool \ where \ ext_check_CSR csrno p isWrite = ( True )\ for csrno :: "(12)Word.word " and p :: " Privilege " and isWrite :: " bool " \ \\ THIS is called if ext_check_CSR returns false. It should cause an appropriate RISCV exception. \\ \ \\val ext_check_CSR_fail : unit -> unit\\ definition ext_check_CSR_fail :: \ unit \ unit \ where \ ext_check_CSR_fail _ = ( () )\ \ \\ Validate a read from physical memory. THIS(access_type, paddr, size, aquire, release, reserved, read_meta) should return Some(exception) to abort the read or None to allow it to proceed. The check is performed after PMP checks and does not apply to MMIO memory. \\ \ \\val ext_check_phys_mem_read : AccessType unit -> mword ty32 -> integer -> bool -> bool -> bool -> bool -> Ext_PhysAddr_Check\\ \ \\ Validate a write to physical memory. THIS(write_kind, paddr, size, data, metadata) should return Some(exception) to abort the write or None to allow it to proceed. The check is performed after PMP checks and does not apply to MMIO memory. \\ \ \\val ext_check_phys_mem_write : forall 'int8_times_n. Size 'int8_times_n => write_kind -> mword ty32 -> integer -> mword 'int8_times_n -> unit -> Ext_PhysAddr_Check\\ \ \\val ext_fetch_check_pc : mword ty32 -> mword ty32 -> Ext_FetchAddr_Check unit\\ definition ext_fetch_check_pc :: \(32)Word.word \(32)Word.word \(unit)Ext_FetchAddr_Check \ where \ ext_fetch_check_pc (start_pc :: xlenbits) (pc :: xlenbits) = ( Ext_FetchAddr_OK pc )\ for start_pc :: "(32)Word.word " and pc :: "(32)Word.word " \ \\val ext_handle_fetch_check_error : unit -> unit\\ definition ext_handle_fetch_check_error :: \ unit \ unit \ where \ ext_handle_fetch_check_error err = ( () )\ for err :: " unit " \ \\val ext_control_check_addr : mword ty32 -> Ext_ControlAddr_Check unit\\ definition ext_control_check_addr :: \(32)Word.word \(unit)Ext_ControlAddr_Check \ where \ ext_control_check_addr pc = ( Ext_ControlAddr_OK pc )\ for pc :: "(32)Word.word " \ \\val ext_control_check_pc : mword ty32 -> Ext_ControlAddr_Check unit\\ definition ext_control_check_pc :: \(32)Word.word \(unit)Ext_ControlAddr_Check \ where \ ext_control_check_pc pc = ( Ext_ControlAddr_OK pc )\ for pc :: "(32)Word.word " \ \\val ext_handle_control_check_error : unit -> unit\\ definition ext_handle_control_check_error :: \ unit \ unit \ where \ ext_handle_control_check_error err = ( () )\ for err :: " unit " \ \\val ext_data_get_addr : mword ty5 -> mword ty32 -> AccessType unit -> word_width -> M (Ext_DataAddr_Check unit)\\ definition ext_data_get_addr :: \(5)Word.word \(32)Word.word \(ext_access_type)AccessType \ word_width \((register_value),((unit)Ext_DataAddr_Check),(exception))monad \ where \ ext_data_get_addr (base :: regidx) (offset :: xlenbits) (acc1 :: ext_access_type AccessType) (width :: word_width) = ( (rX_bits base :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let addr = ((add_vec w__0 offset :: 32 Word.word)) in return (Ext_DataAddr_OK addr)))))\ for base :: "(5)Word.word " and offset :: "(32)Word.word " and acc1 :: "(ext_access_type)AccessType " and width :: " word_width " \ \\val ext_handle_data_check_error : unit -> unit\\ definition ext_handle_data_check_error :: \ unit \ unit \ where \ ext_handle_data_check_error err = ( () )\ for err :: " unit " definition ext_check_phys_mem_read :: \(unit)AccessType \(32)Word.word \ int \ bool \ bool \ bool \ bool \ Ext_PhysAddr_Check \ where \ ext_check_phys_mem_read access_type paddr size1 aquire release reserved read_meta = ( Ext_PhysAddr_OK () )\ for access_type :: "(unit)AccessType " and paddr :: "(32)Word.word " and size1 :: " int " and aquire :: " bool " and release :: " bool " and reserved :: " bool " and read_meta :: " bool " definition ext_check_phys_mem_write :: \ write_kind \(32)Word.word \ int \('int8_times_n::len)Word.word \ unit \ Ext_PhysAddr_Check \ where \ ext_check_phys_mem_write write_kind paddr size1 data metadata = ( Ext_PhysAddr_OK () )\ for write_kind :: " write_kind " and paddr :: "(32)Word.word " and size1 :: " int " and data :: "('int8_times_n::len)Word.word " and metadata :: " unit " \ \\val csr_name_map_forwards : mword ty12 -> string\\ \ \\val csr_name_map_backwards : string -> M (mword ty12)\\ \ \\val csr_name_map_forwards_matches : mword ty12 -> bool\\ \ \\val csr_name_map_backwards_matches : string -> M bool\\ \ \\val csr_name_map_matches_prefix : string -> maybe ((mword ty12 * ii))\\ \ \\val ext_is_CSR_defined : mword ty12 -> Privilege -> M bool\\ \ \\val ext_read_CSR : mword ty12 -> M (maybe (mword ty32))\\ \ \\val ext_write_CSR : mword ty12 -> mword ty32 -> M (maybe (mword ty32))\\ \ \\val undefined_Ustatus : unit -> M Ustatus\\ definition undefined_Ustatus :: \ unit \((register_value),(Ustatus),(exception))monad \ where \ undefined_Ustatus _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Ustatus_bits = w__0 |)))))\ \ \\val Mk_Ustatus : mword ty32 -> Ustatus\\ definition Mk_Ustatus :: \(32)Word.word \ Ustatus \ where \ Mk_Ustatus v = ( (| Ustatus_bits = v |) )\ for v :: "(32)Word.word " definition get_Ustatus_bits :: \ Ustatus \(32)Word.word \ where \ get_Ustatus_bits v = ( (subrange_vec_dec(Ustatus_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Ustatus " definition set_Ustatus_bits :: \((regstate),(register_value),(Ustatus))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Ustatus_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Ustatus_bits := ((update_subrange_vec_dec(Ustatus_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Ustatus))register_ref " and v :: "(32)Word.word " definition update_Ustatus_bits :: \ Ustatus \(32)Word.word \ Ustatus \ where \ update_Ustatus_bits v x = ( ( v (| Ustatus_bits := ((update_subrange_vec_dec(Ustatus_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Ustatus " and x :: "(32)Word.word " definition get_Ustatus_UIE :: \ Ustatus \(1)Word.word \ where \ get_Ustatus_UIE v = ( (subrange_vec_dec(Ustatus_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Ustatus " definition set_Ustatus_UIE :: \((regstate),(register_value),(Ustatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Ustatus_UIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Ustatus_bits := ((update_subrange_vec_dec(Ustatus_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Ustatus))register_ref " and v :: "(1)Word.word " definition update_Ustatus_UIE :: \ Ustatus \(1)Word.word \ Ustatus \ where \ update_Ustatus_UIE v x = ( ( v (| Ustatus_bits := ((update_subrange_vec_dec(Ustatus_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Ustatus " and x :: "(1)Word.word " definition get_Ustatus_UPIE :: \ Ustatus \(1)Word.word \ where \ get_Ustatus_UPIE v = ( (subrange_vec_dec(Ustatus_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Ustatus " definition set_Ustatus_UPIE :: \((regstate),(register_value),(Ustatus))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Ustatus_UPIE r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Ustatus_bits := ((update_subrange_vec_dec(Ustatus_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Ustatus))register_ref " and v :: "(1)Word.word " definition update_Ustatus_UPIE :: \ Ustatus \(1)Word.word \ Ustatus \ where \ update_Ustatus_UPIE v x = ( ( v (| Ustatus_bits := ((update_subrange_vec_dec(Ustatus_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Ustatus " and x :: "(1)Word.word " \ \\val lower_sstatus : Sstatus -> Ustatus\\ definition lower_sstatus :: \ Sstatus \ Ustatus \ where \ lower_sstatus s = ( (let u = (Mk_Ustatus ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) in (let u = (update_Ustatus_UPIE u ((get_Sstatus_UPIE s :: 1 Word.word))) in update_Ustatus_UIE u ((get_Sstatus_UIE s :: 1 Word.word)))))\ for s :: " Sstatus " \ \\val lift_ustatus : Sstatus -> Ustatus -> Sstatus\\ definition lift_ustatus :: \ Sstatus \ Ustatus \ Sstatus \ where \ lift_ustatus (s :: Sstatus) (u :: Ustatus) = ( (let s = (update_Sstatus_UPIE s ((get_Ustatus_UPIE u :: 1 Word.word))) in update_Sstatus_UIE s ((get_Ustatus_UIE u :: 1 Word.word))))\ for s :: " Sstatus " and u :: " Ustatus " \ \\val legalize_ustatus : Mstatus -> mword ty32 -> M Mstatus\\ definition legalize_ustatus :: \ Mstatus \(32)Word.word \((register_value),(Mstatus),(exception))monad \ where \ legalize_ustatus (m :: Mstatus) (v :: xlenbits) = ( (let u = (Mk_Ustatus v) in (let s = (lower_mstatus m) in (let s = (lift_ustatus s u) in lift_sstatus m s))))\ for m :: " Mstatus " and v :: "(32)Word.word " \ \\val undefined_Uinterrupts : unit -> M Uinterrupts\\ definition undefined_Uinterrupts :: \ unit \((register_value),(Uinterrupts),(exception))monad \ where \ undefined_Uinterrupts _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Uinterrupts_bits = w__0 |)))))\ \ \\val Mk_Uinterrupts : mword ty32 -> Uinterrupts\\ definition Mk_Uinterrupts :: \(32)Word.word \ Uinterrupts \ where \ Mk_Uinterrupts v = ( (| Uinterrupts_bits = v |) )\ for v :: "(32)Word.word " definition get_Uinterrupts_bits :: \ Uinterrupts \(32)Word.word \ where \ get_Uinterrupts_bits v = ( (subrange_vec_dec(Uinterrupts_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Uinterrupts " definition set_Uinterrupts_bits :: \((regstate),(register_value),(Uinterrupts))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Uinterrupts_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Uinterrupts_bits := ((update_subrange_vec_dec(Uinterrupts_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Uinterrupts))register_ref " and v :: "(32)Word.word " definition update_Uinterrupts_bits :: \ Uinterrupts \(32)Word.word \ Uinterrupts \ where \ update_Uinterrupts_bits v x = ( ( v (| Uinterrupts_bits := ((update_subrange_vec_dec(Uinterrupts_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Uinterrupts " and x :: "(32)Word.word " definition get_Uinterrupts_UEI :: \ Uinterrupts \(1)Word.word \ where \ get_Uinterrupts_UEI v = ( (subrange_vec_dec(Uinterrupts_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word))\ for v :: " Uinterrupts " definition set_Uinterrupts_UEI :: \((regstate),(register_value),(Uinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Uinterrupts_UEI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Uinterrupts_bits := ((update_subrange_vec_dec(Uinterrupts_bits r) (( 8 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Uinterrupts))register_ref " and v :: "(1)Word.word " definition update_Uinterrupts_UEI :: \ Uinterrupts \(1)Word.word \ Uinterrupts \ where \ update_Uinterrupts_UEI v x = ( ( v (| Uinterrupts_bits := ((update_subrange_vec_dec(Uinterrupts_bits v) (( 8 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Uinterrupts " and x :: "(1)Word.word " definition get_Uinterrupts_USI :: \ Uinterrupts \(1)Word.word \ where \ get_Uinterrupts_USI v = ( (subrange_vec_dec(Uinterrupts_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " Uinterrupts " definition set_Uinterrupts_USI :: \((regstate),(register_value),(Uinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Uinterrupts_USI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Uinterrupts_bits := ((update_subrange_vec_dec(Uinterrupts_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Uinterrupts))register_ref " and v :: "(1)Word.word " definition update_Uinterrupts_USI :: \ Uinterrupts \(1)Word.word \ Uinterrupts \ where \ update_Uinterrupts_USI v x = ( ( v (| Uinterrupts_bits := ((update_subrange_vec_dec(Uinterrupts_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Uinterrupts " and x :: "(1)Word.word " definition get_Uinterrupts_UTI :: \ Uinterrupts \(1)Word.word \ where \ get_Uinterrupts_UTI v = ( (subrange_vec_dec(Uinterrupts_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " Uinterrupts " definition set_Uinterrupts_UTI :: \((regstate),(register_value),(Uinterrupts))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_Uinterrupts_UTI r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Uinterrupts_bits := ((update_subrange_vec_dec(Uinterrupts_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Uinterrupts))register_ref " and v :: "(1)Word.word " definition update_Uinterrupts_UTI :: \ Uinterrupts \(1)Word.word \ Uinterrupts \ where \ update_Uinterrupts_UTI v x = ( ( v (| Uinterrupts_bits := ((update_subrange_vec_dec(Uinterrupts_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Uinterrupts " and x :: "(1)Word.word " \ \\val lower_sip : Sinterrupts -> Sinterrupts -> Uinterrupts\\ definition lower_sip :: \ Sinterrupts \ Sinterrupts \ Uinterrupts \ where \ lower_sip (s :: Sinterrupts) (d :: Sinterrupts) = ( (let (u :: Uinterrupts) = (Mk_Uinterrupts ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) in (let u = (update_Uinterrupts_UEI u ((and_vec ((get_Sinterrupts_UEI s :: 1 Word.word)) ((get_Sinterrupts_UEI d :: 1 Word.word)) :: 1 Word.word))) in (let u = (update_Uinterrupts_UTI u ((and_vec ((get_Sinterrupts_UTI s :: 1 Word.word)) ((get_Sinterrupts_UTI d :: 1 Word.word)) :: 1 Word.word))) in update_Uinterrupts_USI u ((and_vec ((get_Sinterrupts_USI s :: 1 Word.word)) ((get_Sinterrupts_USI d :: 1 Word.word)) :: 1 Word.word))))))\ for s :: " Sinterrupts " and d :: " Sinterrupts " \ \\val lower_sie : Sinterrupts -> Sinterrupts -> Uinterrupts\\ definition lower_sie :: \ Sinterrupts \ Sinterrupts \ Uinterrupts \ where \ lower_sie (s :: Sinterrupts) (d :: Sinterrupts) = ( (let (u :: Uinterrupts) = (Mk_Uinterrupts ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) in (let u = (update_Uinterrupts_UEI u ((and_vec ((get_Sinterrupts_UEI s :: 1 Word.word)) ((get_Sinterrupts_UEI d :: 1 Word.word)) :: 1 Word.word))) in (let u = (update_Uinterrupts_UTI u ((and_vec ((get_Sinterrupts_UTI s :: 1 Word.word)) ((get_Sinterrupts_UTI d :: 1 Word.word)) :: 1 Word.word))) in update_Uinterrupts_USI u ((and_vec ((get_Sinterrupts_USI s :: 1 Word.word)) ((get_Sinterrupts_USI d :: 1 Word.word)) :: 1 Word.word))))))\ for s :: " Sinterrupts " and d :: " Sinterrupts " \ \\val lift_uip : Sinterrupts -> Sinterrupts -> Uinterrupts -> Sinterrupts\\ definition lift_uip :: \ Sinterrupts \ Sinterrupts \ Uinterrupts \ Sinterrupts \ where \ lift_uip (o1 :: Sinterrupts) (d :: Sinterrupts) (u :: Uinterrupts) = ( (let (s :: Sinterrupts) = o1 in if (((((get_Sinterrupts_USI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Sinterrupts_USI s ((get_Uinterrupts_USI u :: 1 Word.word)) else s))\ for o1 :: " Sinterrupts " and d :: " Sinterrupts " and u :: " Uinterrupts " \ \\val legalize_uip : Sinterrupts -> Sinterrupts -> mword ty32 -> Sinterrupts\\ definition legalize_uip :: \ Sinterrupts \ Sinterrupts \(32)Word.word \ Sinterrupts \ where \ legalize_uip (s :: Sinterrupts) (d :: Sinterrupts) (v :: xlenbits) = ( lift_uip s d ((Mk_Uinterrupts v)))\ for s :: " Sinterrupts " and d :: " Sinterrupts " and v :: "(32)Word.word " \ \\val lift_uie : Sinterrupts -> Sinterrupts -> Uinterrupts -> Sinterrupts\\ definition lift_uie :: \ Sinterrupts \ Sinterrupts \ Uinterrupts \ Sinterrupts \ where \ lift_uie (o1 :: Sinterrupts) (d :: Sinterrupts) (u :: Uinterrupts) = ( (let (s :: Sinterrupts) = o1 in (let s = (if (((((get_Sinterrupts_UEI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Sinterrupts_UEI s ((get_Uinterrupts_UEI u :: 1 Word.word)) else s) in (let s = (if (((((get_Sinterrupts_UTI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Sinterrupts_UTI s ((get_Uinterrupts_UTI u :: 1 Word.word)) else s) in if (((((get_Sinterrupts_USI d :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then update_Sinterrupts_USI s ((get_Uinterrupts_USI u :: 1 Word.word)) else s))))\ for o1 :: " Sinterrupts " and d :: " Sinterrupts " and u :: " Uinterrupts " \ \\val legalize_uie : Sinterrupts -> Sinterrupts -> mword ty32 -> Sinterrupts\\ definition legalize_uie :: \ Sinterrupts \ Sinterrupts \(32)Word.word \ Sinterrupts \ where \ legalize_uie (s :: Sinterrupts) (d :: Sinterrupts) (v :: xlenbits) = ( lift_uie s d ((Mk_Uinterrupts v)))\ for s :: " Sinterrupts " and d :: " Sinterrupts " and v :: "(32)Word.word " \ \\val ext_check_xret_priv : Privilege -> bool\\ definition ext_check_xret_priv :: \ Privilege \ bool \ where \ ext_check_xret_priv p = ( True )\ for p :: " Privilege " \ \\val ext_fail_xret_priv : unit -> unit\\ definition ext_fail_xret_priv :: \ unit \ unit \ where \ ext_fail_xret_priv _ = ( () )\ \ \\val handle_trap_extension : Privilege -> mword ty32 -> maybe unit -> unit\\ definition handle_trap_extension :: \ Privilege \(32)Word.word \(unit)option \ unit \ where \ handle_trap_extension (p :: Privilege) (pc :: xlenbits) (u :: unit option) = ( () )\ for p :: " Privilege " and pc :: "(32)Word.word " and u :: "(unit)option " \ \\val prepare_trap_vector : Privilege -> Mcause -> M (mword ty32)\\ definition prepare_trap_vector :: \ Privilege \ Mcause \((register_value),((32)Word.word),(exception))monad \ where \ prepare_trap_vector (p :: Privilege) (cause :: Mcause) = ( (case p of Machine => read_reg mtvec_ref | Supervisor => read_reg stvec_ref | User => read_reg utvec_ref ) \ ((\ (tvec :: Mtvec) . (case ((tvec_addr tvec cause :: ( 32 Word.word)option)) of Some (epc) => return epc | None => (internal_error (''Invalid tvec mode'') :: ( 32 Word.word) M) ))))\ for p :: " Privilege " and cause :: " Mcause " \ \\val get_xret_target : Privilege -> M (mword ty32)\\ fun get_xret_target :: \ Privilege \((register_value),((32)Word.word),(exception))monad \ where \ get_xret_target Machine = ( (read_reg mepc_ref :: ( 32 Word.word) M))\ |\ get_xret_target Supervisor = ( (read_reg sepc_ref :: ( 32 Word.word) M))\ |\ get_xret_target User = ( (read_reg uepc_ref :: ( 32 Word.word) M))\ \ \\val set_xret_target : Privilege -> mword ty32 -> M (mword ty32)\\ definition set_xret_target :: \ Privilege \(32)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ set_xret_target p value1 = ( (legalize_xepc value1 :: ( 32 Word.word) M) \ ((\ target . (case p of Machine => write_reg mepc_ref target | Supervisor => write_reg sepc_ref target | User => write_reg uepc_ref target ) \ return target)))\ for p :: " Privilege " and value1 :: "(32)Word.word " \ \\val prepare_xret_target : Privilege -> M (mword ty32)\\ definition prepare_xret_target :: \ Privilege \((register_value),((32)Word.word),(exception))monad \ where \ prepare_xret_target p = ( (get_xret_target p :: ( 32 Word.word) M))\ for p :: " Privilege " \ \\val get_mtvec : unit -> M (mword ty32)\\ definition get_mtvec :: \ unit \((register_value),((32)Word.word),(exception))monad \ where \ get_mtvec _ = ( read_reg mtvec_ref \ ((\ (w__0 :: Mtvec) . return ((get_Mtvec_bits w__0 :: 32 Word.word)))))\ \ \\val get_stvec : unit -> M (mword ty32)\\ definition get_stvec :: \ unit \((register_value),((32)Word.word),(exception))monad \ where \ get_stvec _ = ( read_reg stvec_ref \ ((\ (w__0 :: Mtvec) . return ((get_Mtvec_bits w__0 :: 32 Word.word)))))\ \ \\val get_utvec : unit -> M (mword ty32)\\ definition get_utvec :: \ unit \((register_value),((32)Word.word),(exception))monad \ where \ get_utvec _ = ( read_reg utvec_ref \ ((\ (w__0 :: Mtvec) . return ((get_Mtvec_bits w__0 :: 32 Word.word)))))\ \ \\val set_mtvec : mword ty32 -> M (mword ty32)\\ definition set_mtvec :: \(32)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ set_mtvec value1 = ( read_reg mtvec_ref \ ((\ (w__0 :: Mtvec) . (write_reg mtvec_ref ((legalize_tvec w__0 value1)) \ read_reg mtvec_ref) \ ((\ (w__1 :: Mtvec) . return ((get_Mtvec_bits w__1 :: 32 Word.word)))))))\ for value1 :: "(32)Word.word " \ \\val set_stvec : mword ty32 -> M (mword ty32)\\ definition set_stvec :: \(32)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ set_stvec value1 = ( read_reg stvec_ref \ ((\ (w__0 :: Mtvec) . (write_reg stvec_ref ((legalize_tvec w__0 value1)) \ read_reg stvec_ref) \ ((\ (w__1 :: Mtvec) . return ((get_Mtvec_bits w__1 :: 32 Word.word)))))))\ for value1 :: "(32)Word.word " \ \\val set_utvec : mword ty32 -> M (mword ty32)\\ definition set_utvec :: \(32)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ set_utvec value1 = ( read_reg utvec_ref \ ((\ (w__0 :: Mtvec) . (write_reg utvec_ref ((legalize_tvec w__0 value1)) \ read_reg utvec_ref) \ ((\ (w__1 :: Mtvec) . return ((get_Mtvec_bits w__1 :: 32 Word.word)))))))\ for value1 :: "(32)Word.word " \ \\val update_softfloat_fflags : mword ty5 -> M unit\\ definition update_softfloat_fflags :: \(5)Word.word \((register_value),(unit),(exception))monad \ where \ update_softfloat_fflags flags = ( write_reg float_fflags_ref ((zero_extend flags (( 64 :: int)::ii) :: 64 Word.word)))\ for flags :: "(5)Word.word " \ \\val riscv_f16Add : mword ty3 -> mword ty16 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16Add :: \(3)Word.word \(16)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16Add rm v1 v2 = ( (let (_ :: unit) = (softfloat_f16_add rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(16)Word.word " and v2 :: "(16)Word.word " \ \\val riscv_f16Sub : mword ty3 -> mword ty16 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16Sub :: \(3)Word.word \(16)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16Sub rm v1 v2 = ( (let (_ :: unit) = (softfloat_f16_sub rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(16)Word.word " and v2 :: "(16)Word.word " \ \\val riscv_f16Mul : mword ty3 -> mword ty16 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16Mul :: \(3)Word.word \(16)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16Mul rm v1 v2 = ( (let (_ :: unit) = (softfloat_f16_mul rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(16)Word.word " and v2 :: "(16)Word.word " \ \\val riscv_f16Div : mword ty3 -> mword ty16 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16Div :: \(3)Word.word \(16)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16Div rm v1 v2 = ( (let (_ :: unit) = (softfloat_f16_div rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(16)Word.word " and v2 :: "(16)Word.word " \ \\val riscv_f32Add : mword ty3 -> mword ty32 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32Add :: \(3)Word.word \(32)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32Add rm v1 v2 = ( (let (_ :: unit) = (softfloat_f32_add rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(32)Word.word " and v2 :: "(32)Word.word " \ \\val riscv_f32Sub : mword ty3 -> mword ty32 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32Sub :: \(3)Word.word \(32)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32Sub rm v1 v2 = ( (let (_ :: unit) = (softfloat_f32_sub rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(32)Word.word " and v2 :: "(32)Word.word " \ \\val riscv_f32Mul : mword ty3 -> mword ty32 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32Mul :: \(3)Word.word \(32)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32Mul rm v1 v2 = ( (let (_ :: unit) = (softfloat_f32_mul rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(32)Word.word " and v2 :: "(32)Word.word " \ \\val riscv_f32Div : mword ty3 -> mword ty32 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32Div :: \(3)Word.word \(32)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32Div rm v1 v2 = ( (let (_ :: unit) = (softfloat_f32_div rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(32)Word.word " and v2 :: "(32)Word.word " \ \\val riscv_f64Add : mword ty3 -> mword ty64 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64Add :: \(3)Word.word \(64)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64Add rm v1 v2 = ( (let (_ :: unit) = (softfloat_f64_add rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v1 :: "(64)Word.word " and v2 :: "(64)Word.word " \ \\val riscv_f64Sub : mword ty3 -> mword ty64 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64Sub :: \(3)Word.word \(64)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64Sub rm v1 v2 = ( (let (_ :: unit) = (softfloat_f64_sub rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v1 :: "(64)Word.word " and v2 :: "(64)Word.word " \ \\val riscv_f64Mul : mword ty3 -> mword ty64 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64Mul :: \(3)Word.word \(64)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64Mul rm v1 v2 = ( (let (_ :: unit) = (softfloat_f64_mul rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v1 :: "(64)Word.word " and v2 :: "(64)Word.word " \ \\val riscv_f64Div : mword ty3 -> mword ty64 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64Div :: \(3)Word.word \(64)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64Div rm v1 v2 = ( (let (_ :: unit) = (softfloat_f64_div rm v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v1 :: "(64)Word.word " and v2 :: "(64)Word.word " \ \\val riscv_f16MulAdd : mword ty3 -> mword ty16 -> mword ty16 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16MulAdd :: \(3)Word.word \(16)Word.word \(16)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16MulAdd rm v1 v2 v3 = ( (let (_ :: unit) = (softfloat_f16_muladd rm v1 v2 v3) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(16)Word.word " and v2 :: "(16)Word.word " and v3 :: "(16)Word.word " \ \\val riscv_f32MulAdd : mword ty3 -> mword ty32 -> mword ty32 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32MulAdd :: \(3)Word.word \(32)Word.word \(32)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32MulAdd rm v1 v2 v3 = ( (let (_ :: unit) = (softfloat_f32_muladd rm v1 v2 v3) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v1 :: "(32)Word.word " and v2 :: "(32)Word.word " and v3 :: "(32)Word.word " \ \\val riscv_f64MulAdd : mword ty3 -> mword ty64 -> mword ty64 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64MulAdd :: \(3)Word.word \(64)Word.word \(64)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64MulAdd rm v1 v2 v3 = ( (let (_ :: unit) = (softfloat_f64_muladd rm v1 v2 v3) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v1 :: "(64)Word.word " and v2 :: "(64)Word.word " and v3 :: "(64)Word.word " \ \\val riscv_f16Sqrt : mword ty3 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16Sqrt :: \(3)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16Sqrt rm v = ( (let (_ :: unit) = (softfloat_f16_sqrt rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(16)Word.word " \ \\val riscv_f32Sqrt : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32Sqrt :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32Sqrt rm v = ( (let (_ :: unit) = (softfloat_f32_sqrt rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_f64Sqrt : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64Sqrt :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64Sqrt rm v = ( (let (_ :: unit) = (softfloat_f64_sqrt rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_f16ToI32 : mword ty3 -> mword ty16 -> M (mword ty5 * mword ty32)\\ definition riscv_f16ToI32 :: \(3)Word.word \(16)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f16ToI32 rm v = ( (let (_ :: unit) = (softfloat_f16_to_i32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(16)Word.word " \ \\val riscv_f16ToUi32 : mword ty3 -> mword ty16 -> M (mword ty5 * mword ty32)\\ definition riscv_f16ToUi32 :: \(3)Word.word \(16)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f16ToUi32 rm v = ( (let (_ :: unit) = (softfloat_f16_to_ui32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(16)Word.word " \ \\val riscv_i32ToF16 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty16)\\ definition riscv_i32ToF16 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_i32ToF16 rm v = ( (let (_ :: unit) = (softfloat_i32_to_f16 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_ui32ToF16 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty16)\\ definition riscv_ui32ToF16 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_ui32ToF16 rm v = ( (let (_ :: unit) = (softfloat_ui32_to_f16 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_f16ToI64 : mword ty3 -> mword ty16 -> M (mword ty5 * mword ty64)\\ definition riscv_f16ToI64 :: \(3)Word.word \(16)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f16ToI64 rm v = ( (let (_ :: unit) = (softfloat_f16_to_i64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(16)Word.word " \ \\val riscv_f16ToUi64 : mword ty3 -> mword ty16 -> M (mword ty5 * mword ty64)\\ definition riscv_f16ToUi64 :: \(3)Word.word \(16)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f16ToUi64 rm v = ( (let (_ :: unit) = (softfloat_f16_to_ui64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(16)Word.word " \ \\val riscv_i64ToF16 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty16)\\ definition riscv_i64ToF16 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_i64ToF16 rm v = ( (let (_ :: unit) = (softfloat_i64_to_f16 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_ui64ToF16 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty16)\\ definition riscv_ui64ToF16 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_ui64ToF16 rm v = ( (let (_ :: unit) = (softfloat_ui64_to_f16 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_f32ToI32 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32ToI32 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32ToI32 rm v = ( (let (_ :: unit) = (softfloat_f32_to_i32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_f32ToUi32 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32ToUi32 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32ToUi32 rm v = ( (let (_ :: unit) = (softfloat_f32_to_ui32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_i32ToF32 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_i32ToF32 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_i32ToF32 rm v = ( (let (_ :: unit) = (softfloat_i32_to_f32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_ui32ToF32 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_ui32ToF32 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_ui32ToF32 rm v = ( (let (_ :: unit) = (softfloat_ui32_to_f32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_f32ToI64 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty64)\\ definition riscv_f32ToI64 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f32ToI64 rm v = ( (let (_ :: unit) = (softfloat_f32_to_i64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_f32ToUi64 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty64)\\ definition riscv_f32ToUi64 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f32ToUi64 rm v = ( (let (_ :: unit) = (softfloat_f32_to_ui64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_i64ToF32 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty32)\\ definition riscv_i64ToF32 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_i64ToF32 rm v = ( (let (_ :: unit) = (softfloat_i64_to_f32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_ui64ToF32 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty32)\\ definition riscv_ui64ToF32 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_ui64ToF32 rm v = ( (let (_ :: unit) = (softfloat_ui64_to_f32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_f64ToI32 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty32)\\ definition riscv_f64ToI32 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f64ToI32 rm v = ( (let (_ :: unit) = (softfloat_f64_to_i32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_f64ToUi32 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty32)\\ definition riscv_f64ToUi32 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f64ToUi32 rm v = ( (let (_ :: unit) = (softfloat_f64_to_ui32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_i32ToF64 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty64)\\ definition riscv_i32ToF64 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_i32ToF64 rm v = ( (let (_ :: unit) = (softfloat_i32_to_f64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_ui32ToF64 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty64)\\ definition riscv_ui32ToF64 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_ui32ToF64 rm v = ( (let (_ :: unit) = (softfloat_ui32_to_f64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_f64ToI64 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64ToI64 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64ToI64 rm v = ( (let (_ :: unit) = (softfloat_f64_to_i64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_f64ToUi64 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64ToUi64 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64ToUi64 rm v = ( (let (_ :: unit) = (softfloat_f64_to_ui64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_i64ToF64 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_i64ToF64 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_i64ToF64 rm v = ( (let (_ :: unit) = (softfloat_i64_to_f64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_ui64ToF64 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_ui64ToF64 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_ui64ToF64 rm v = ( (let (_ :: unit) = (softfloat_ui64_to_f64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_f16ToF32 : mword ty3 -> mword ty16 -> M (mword ty5 * mword ty32)\\ definition riscv_f16ToF32 :: \(3)Word.word \(16)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f16ToF32 rm v = ( (let (_ :: unit) = (softfloat_f16_to_f32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(16)Word.word " \ \\val riscv_f16ToF64 : mword ty3 -> mword ty16 -> M (mword ty5 * mword ty64)\\ definition riscv_f16ToF64 :: \(3)Word.word \(16)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f16ToF64 rm v = ( (let (_ :: unit) = (softfloat_f16_to_f64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(16)Word.word " \ \\val riscv_f32ToF64 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty64)\\ definition riscv_f32ToF64 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f32ToF64 rm v = ( (let (_ :: unit) = (softfloat_f32_to_f64 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_f32ToF16 : mword ty3 -> mword ty32 -> M (mword ty5 * mword ty16)\\ definition riscv_f32ToF16 :: \(3)Word.word \(32)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f32ToF16 rm v = ( (let (_ :: unit) = (softfloat_f32_to_f16 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(32)Word.word " \ \\val riscv_f64ToF16 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty16)\\ definition riscv_f64ToF16 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f64ToF16 rm v = ( (let (_ :: unit) = (softfloat_f64_to_f16 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_f64ToF32 : mword ty3 -> mword ty64 -> M (mword ty5 * mword ty32)\\ definition riscv_f64ToF32 :: \(3)Word.word \(64)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f64ToF32 rm v = ( (let (_ :: unit) = (softfloat_f64_to_f32 rm v) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for rm :: "(3)Word.word " and v :: "(64)Word.word " \ \\val riscv_f16Lt : mword ty16 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16Lt :: \(16)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16Lt v1 v2 = ( (let (_ :: unit) = (softfloat_f16_lt v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for v1 :: "(16)Word.word " and v2 :: "(16)Word.word " \ \\val riscv_f16Le : mword ty16 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16Le :: \(16)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16Le v1 v2 = ( (let (_ :: unit) = (softfloat_f16_le v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for v1 :: "(16)Word.word " and v2 :: "(16)Word.word " \ \\val riscv_f16Eq : mword ty16 -> mword ty16 -> M (mword ty5 * mword ty16)\\ definition riscv_f16Eq :: \(16)Word.word \(16)Word.word \((register_value),((5)Word.word*(16)Word.word),(exception))monad \ where \ riscv_f16Eq v1 v2 = ( (let (_ :: unit) = (softfloat_f16_eq v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word))))))))\ for v1 :: "(16)Word.word " and v2 :: "(16)Word.word " \ \\val riscv_f32Lt : mword ty32 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32Lt :: \(32)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32Lt v1 v2 = ( (let (_ :: unit) = (softfloat_f32_lt v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for v1 :: "(32)Word.word " and v2 :: "(32)Word.word " \ \\val riscv_f32Le : mword ty32 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32Le :: \(32)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32Le v1 v2 = ( (let (_ :: unit) = (softfloat_f32_le v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for v1 :: "(32)Word.word " and v2 :: "(32)Word.word " \ \\val riscv_f32Eq : mword ty32 -> mword ty32 -> M (mword ty5 * mword ty32)\\ definition riscv_f32Eq :: \(32)Word.word \(32)Word.word \((register_value),((5)Word.word*(32)Word.word),(exception))monad \ where \ riscv_f32Eq v1 v2 = ( (let (_ :: unit) = (softfloat_f32_eq v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), (subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))))))\ for v1 :: "(32)Word.word " and v2 :: "(32)Word.word " \ \\val riscv_f64Lt : mword ty64 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64Lt :: \(64)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64Lt v1 v2 = ( (let (_ :: unit) = (softfloat_f64_lt v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for v1 :: "(64)Word.word " and v2 :: "(64)Word.word " \ \\val riscv_f64Le : mword ty64 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64Le :: \(64)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64Le v1 v2 = ( (let (_ :: unit) = (softfloat_f64_le v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for v1 :: "(64)Word.word " and v2 :: "(64)Word.word " \ \\val riscv_f64Eq : mword ty64 -> mword ty64 -> M (mword ty5 * mword ty64)\\ definition riscv_f64Eq :: \(64)Word.word \(64)Word.word \((register_value),((5)Word.word*(64)Word.word),(exception))monad \ where \ riscv_f64Eq v1 v2 = ( (let (_ :: unit) = (softfloat_f64_eq v1 v2) in (read_reg float_fflags_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg float_result_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return ((subrange_vec_dec w__0 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word), w__1)))))))\ for v1 :: "(64)Word.word " and v2 :: "(64)Word.word " \ \\val dirty_fd_context : unit -> M unit\\ definition dirty_fd_context :: \ unit \((register_value),(unit),(exception))monad \ where \ dirty_fd_context _ = ( set_Mstatus_FS mstatus_ref ((extStatus_to_bits Dirty :: 2 Word.word)) \ set_Mstatus_SD mstatus_ref ( 0b1 :: 1 Word.word))\ \ \\val rF : integer -> M (mword ty32)\\ definition rF :: \ int \((register_value),((32)Word.word),(exception))monad \ where \ rF r = ( (let l__45 = r in (if (((l__45 = (( 0 :: int)::ii)))) then (read_reg f0_ref :: ( 32 Word.word) M) else if (((l__45 = (( 1 :: int)::ii)))) then (read_reg f1_ref :: ( 32 Word.word) M) else if (((l__45 = (( 2 :: int)::ii)))) then (read_reg f2_ref :: ( 32 Word.word) M) else if (((l__45 = (( 3 :: int)::ii)))) then (read_reg f3_ref :: ( 32 Word.word) M) else if (((l__45 = (( 4 :: int)::ii)))) then (read_reg f4_ref :: ( 32 Word.word) M) else if (((l__45 = (( 5 :: int)::ii)))) then (read_reg f5_ref :: ( 32 Word.word) M) else if (((l__45 = (( 6 :: int)::ii)))) then (read_reg f6_ref :: ( 32 Word.word) M) else if (((l__45 = (( 7 :: int)::ii)))) then (read_reg f7_ref :: ( 32 Word.word) M) else if (((l__45 = (( 8 :: int)::ii)))) then (read_reg f8_ref :: ( 32 Word.word) M) else if (((l__45 = (( 9 :: int)::ii)))) then (read_reg f9_ref :: ( 32 Word.word) M) else if (((l__45 = (( 10 :: int)::ii)))) then (read_reg f10_ref :: ( 32 Word.word) M) else if (((l__45 = (( 11 :: int)::ii)))) then (read_reg f11_ref :: ( 32 Word.word) M) else if (((l__45 = (( 12 :: int)::ii)))) then (read_reg f12_ref :: ( 32 Word.word) M) else if (((l__45 = (( 13 :: int)::ii)))) then (read_reg f13_ref :: ( 32 Word.word) M) else if (((l__45 = (( 14 :: int)::ii)))) then (read_reg f14_ref :: ( 32 Word.word) M) else if (((l__45 = (( 15 :: int)::ii)))) then (read_reg f15_ref :: ( 32 Word.word) M) else if (((l__45 = (( 16 :: int)::ii)))) then (read_reg f16_ref :: ( 32 Word.word) M) else if (((l__45 = (( 17 :: int)::ii)))) then (read_reg f17_ref :: ( 32 Word.word) M) else if (((l__45 = (( 18 :: int)::ii)))) then (read_reg f18_ref :: ( 32 Word.word) M) else if (((l__45 = (( 19 :: int)::ii)))) then (read_reg f19_ref :: ( 32 Word.word) M) else if (((l__45 = (( 20 :: int)::ii)))) then (read_reg f20_ref :: ( 32 Word.word) M) else if (((l__45 = (( 21 :: int)::ii)))) then (read_reg f21_ref :: ( 32 Word.word) M) else if (((l__45 = (( 22 :: int)::ii)))) then (read_reg f22_ref :: ( 32 Word.word) M) else if (((l__45 = (( 23 :: int)::ii)))) then (read_reg f23_ref :: ( 32 Word.word) M) else if (((l__45 = (( 24 :: int)::ii)))) then (read_reg f24_ref :: ( 32 Word.word) M) else if (((l__45 = (( 25 :: int)::ii)))) then (read_reg f25_ref :: ( 32 Word.word) M) else if (((l__45 = (( 26 :: int)::ii)))) then (read_reg f26_ref :: ( 32 Word.word) M) else if (((l__45 = (( 27 :: int)::ii)))) then (read_reg f27_ref :: ( 32 Word.word) M) else if (((l__45 = (( 28 :: int)::ii)))) then (read_reg f28_ref :: ( 32 Word.word) M) else if (((l__45 = (( 29 :: int)::ii)))) then (read_reg f29_ref :: ( 32 Word.word) M) else if (((l__45 = (( 30 :: int)::ii)))) then (read_reg f30_ref :: ( 32 Word.word) M) else if (((l__45 = (( 31 :: int)::ii)))) then (read_reg f31_ref :: ( 32 Word.word) M) else assert_exp False (''invalid floating point register number'') \ exit0 () ) \ ((\ (v :: fregtype) . return ((fregval_from_freg v :: 32 Word.word))))))\ for r :: " int " \ \\val wF : integer -> mword ty32 -> M unit\\ definition wF :: \ int \(32)Word.word \((register_value),(unit),(exception))monad \ where \ wF r in_v = ( (let v = ((fregval_into_freg in_v :: 32 Word.word)) in (let l__13 = r in ((if (((l__13 = (( 0 :: int)::ii)))) then write_reg f0_ref v else if (((l__13 = (( 1 :: int)::ii)))) then write_reg f1_ref v else if (((l__13 = (( 2 :: int)::ii)))) then write_reg f2_ref v else if (((l__13 = (( 3 :: int)::ii)))) then write_reg f3_ref v else if (((l__13 = (( 4 :: int)::ii)))) then write_reg f4_ref v else if (((l__13 = (( 5 :: int)::ii)))) then write_reg f5_ref v else if (((l__13 = (( 6 :: int)::ii)))) then write_reg f6_ref v else if (((l__13 = (( 7 :: int)::ii)))) then write_reg f7_ref v else if (((l__13 = (( 8 :: int)::ii)))) then write_reg f8_ref v else if (((l__13 = (( 9 :: int)::ii)))) then write_reg f9_ref v else if (((l__13 = (( 10 :: int)::ii)))) then write_reg f10_ref v else if (((l__13 = (( 11 :: int)::ii)))) then write_reg f11_ref v else if (((l__13 = (( 12 :: int)::ii)))) then write_reg f12_ref v else if (((l__13 = (( 13 :: int)::ii)))) then write_reg f13_ref v else if (((l__13 = (( 14 :: int)::ii)))) then write_reg f14_ref v else if (((l__13 = (( 15 :: int)::ii)))) then write_reg f15_ref v else if (((l__13 = (( 16 :: int)::ii)))) then write_reg f16_ref v else if (((l__13 = (( 17 :: int)::ii)))) then write_reg f17_ref v else if (((l__13 = (( 18 :: int)::ii)))) then write_reg f18_ref v else if (((l__13 = (( 19 :: int)::ii)))) then write_reg f19_ref v else if (((l__13 = (( 20 :: int)::ii)))) then write_reg f20_ref v else if (((l__13 = (( 21 :: int)::ii)))) then write_reg f21_ref v else if (((l__13 = (( 22 :: int)::ii)))) then write_reg f22_ref v else if (((l__13 = (( 23 :: int)::ii)))) then write_reg f23_ref v else if (((l__13 = (( 24 :: int)::ii)))) then write_reg f24_ref v else if (((l__13 = (( 25 :: int)::ii)))) then write_reg f25_ref v else if (((l__13 = (( 26 :: int)::ii)))) then write_reg f26_ref v else if (((l__13 = (( 27 :: int)::ii)))) then write_reg f27_ref v else if (((l__13 = (( 28 :: int)::ii)))) then write_reg f28_ref v else if (((l__13 = (( 29 :: int)::ii)))) then write_reg f29_ref v else if (((l__13 = (( 30 :: int)::ii)))) then write_reg f30_ref v else if (((l__13 = (( 31 :: int)::ii)))) then write_reg f31_ref v else assert_exp False (''invalid floating point register number'') \ exit0 () ) \ dirty_fd_context () ) \ return (if ((get_config_print_reg () )) then print_dbg (((@) (''f'') (((@) ((stringFromInteger r)) (((@) ('' <- '') ((FRegStr v)))))))) else () ))))\ for r :: " int " and in_v :: "(32)Word.word " \ \\val rF_bits : mword ty5 -> M (mword ty32)\\ definition rF_bits :: \(5)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ rF_bits i = ( (rF ((Word.uint i)) :: ( 32 Word.word) M))\ for i :: "(5)Word.word " \ \\val wF_bits : mword ty5 -> mword ty32 -> M unit\\ definition wF_bits :: \(5)Word.word \(32)Word.word \((register_value),(unit),(exception))monad \ where \ wF_bits (i :: 5 bits) (data :: flenbits) = ( wF ((Word.uint i)) data )\ for i :: "(5)Word.word " and data :: "(32)Word.word " \ \\val freg_name_abi_forwards : mword ty5 -> M string\\ \ \\val freg_name_abi_backwards : string -> M (mword ty5)\\ \ \\val freg_name_abi_forwards_matches : mword ty5 -> bool\\ \ \\val freg_name_abi_backwards_matches : string -> bool\\ \ \\val freg_name_abi_matches_prefix : string -> maybe ((mword ty5 * ii))\\ definition freg_name_abi_forwards :: \(5)Word.word \((register_value),(string),(exception))monad \ where \ freg_name_abi_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00000 :: 5 Word.word)))) then return (''ft0'') else if (((b__0 = ( 0b00001 :: 5 Word.word)))) then return (''ft1'') else if (((b__0 = ( 0b00010 :: 5 Word.word)))) then return (''ft2'') else if (((b__0 = ( 0b00011 :: 5 Word.word)))) then return (''ft3'') else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then return (''ft4'') else if (((b__0 = ( 0b00101 :: 5 Word.word)))) then return (''ft5'') else if (((b__0 = ( 0b00110 :: 5 Word.word)))) then return (''ft6'') else if (((b__0 = ( 0b00111 :: 5 Word.word)))) then return (''ft7'') else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then return (''fs0'') else if (((b__0 = ( 0b01001 :: 5 Word.word)))) then return (''fs1'') else if (((b__0 = ( 0b01010 :: 5 Word.word)))) then return (''fa0'') else if (((b__0 = ( 0b01011 :: 5 Word.word)))) then return (''fa1'') else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then return (''fa2'') else if (((b__0 = ( 0b01101 :: 5 Word.word)))) then return (''fa3'') else if (((b__0 = ( 0b01110 :: 5 Word.word)))) then return (''fa4'') else if (((b__0 = ( 0b01111 :: 5 Word.word)))) then return (''fa5'') else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then return (''fa6'') else if (((b__0 = ( 0b10001 :: 5 Word.word)))) then return (''fa7'') else if (((b__0 = ( 0b10010 :: 5 Word.word)))) then return (''fs2'') else if (((b__0 = ( 0b10011 :: 5 Word.word)))) then return (''fs3'') else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then return (''fs4'') else if (((b__0 = ( 0b10101 :: 5 Word.word)))) then return (''fs5'') else if (((b__0 = ( 0b10110 :: 5 Word.word)))) then return (''fs6'') else if (((b__0 = ( 0b10111 :: 5 Word.word)))) then return (''fs7'') else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then return (''fs8'') else if (((b__0 = ( 0b11001 :: 5 Word.word)))) then return (''fs9'') else if (((b__0 = ( 0b11010 :: 5 Word.word)))) then return (''fs10'') else if (((b__0 = ( 0b11011 :: 5 Word.word)))) then return (''fs11'') else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then return (''ft8'') else if (((b__0 = ( 0b11101 :: 5 Word.word)))) then return (''ft9'') else if (((b__0 = ( 0b11110 :: 5 Word.word)))) then return (''ft10'') else if (((b__0 = ( 0b11111 :: 5 Word.word)))) then return (''ft11'') else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(5)Word.word " definition freg_name_abi_backwards :: \ string \((register_value),((5)Word.word),(exception))monad \ where \ freg_name_abi_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''ft0'')))) then return ( 0b00000 :: 5 Word.word) else if (((p00 = (''ft1'')))) then return ( 0b00001 :: 5 Word.word) else if (((p00 = (''ft2'')))) then return ( 0b00010 :: 5 Word.word) else if (((p00 = (''ft3'')))) then return ( 0b00011 :: 5 Word.word) else if (((p00 = (''ft4'')))) then return ( 0b00100 :: 5 Word.word) else if (((p00 = (''ft5'')))) then return ( 0b00101 :: 5 Word.word) else if (((p00 = (''ft6'')))) then return ( 0b00110 :: 5 Word.word) else if (((p00 = (''ft7'')))) then return ( 0b00111 :: 5 Word.word) else if (((p00 = (''fs0'')))) then return ( 0b01000 :: 5 Word.word) else if (((p00 = (''fs1'')))) then return ( 0b01001 :: 5 Word.word) else if (((p00 = (''fa0'')))) then return ( 0b01010 :: 5 Word.word) else if (((p00 = (''fa1'')))) then return ( 0b01011 :: 5 Word.word) else if (((p00 = (''fa2'')))) then return ( 0b01100 :: 5 Word.word) else if (((p00 = (''fa3'')))) then return ( 0b01101 :: 5 Word.word) else if (((p00 = (''fa4'')))) then return ( 0b01110 :: 5 Word.word) else if (((p00 = (''fa5'')))) then return ( 0b01111 :: 5 Word.word) else if (((p00 = (''fa6'')))) then return ( 0b10000 :: 5 Word.word) else if (((p00 = (''fa7'')))) then return ( 0b10001 :: 5 Word.word) else if (((p00 = (''fs2'')))) then return ( 0b10010 :: 5 Word.word) else if (((p00 = (''fs3'')))) then return ( 0b10011 :: 5 Word.word) else if (((p00 = (''fs4'')))) then return ( 0b10100 :: 5 Word.word) else if (((p00 = (''fs5'')))) then return ( 0b10101 :: 5 Word.word) else if (((p00 = (''fs6'')))) then return ( 0b10110 :: 5 Word.word) else if (((p00 = (''fs7'')))) then return ( 0b10111 :: 5 Word.word) else if (((p00 = (''fs8'')))) then return ( 0b11000 :: 5 Word.word) else if (((p00 = (''fs9'')))) then return ( 0b11001 :: 5 Word.word) else if (((p00 = (''fs10'')))) then return ( 0b11010 :: 5 Word.word) else if (((p00 = (''fs11'')))) then return ( 0b11011 :: 5 Word.word) else if (((p00 = (''ft8'')))) then return ( 0b11100 :: 5 Word.word) else if (((p00 = (''ft9'')))) then return ( 0b11101 :: 5 Word.word) else if (((p00 = (''ft10'')))) then return ( 0b11110 :: 5 Word.word) else if (((p00 = (''ft11'')))) then return ( 0b11111 :: 5 Word.word) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition freg_name_abi_forwards_matches :: \(5)Word.word \ bool \ where \ freg_name_abi_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11111 :: 5 Word.word)))) then True else False))\ for arg1 :: "(5)Word.word " definition freg_name_abi_backwards_matches :: \ string \ bool \ where \ freg_name_abi_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''ft0'')))) then True else if (((p00 = (''ft1'')))) then True else if (((p00 = (''ft2'')))) then True else if (((p00 = (''ft3'')))) then True else if (((p00 = (''ft4'')))) then True else if (((p00 = (''ft5'')))) then True else if (((p00 = (''ft6'')))) then True else if (((p00 = (''ft7'')))) then True else if (((p00 = (''fs0'')))) then True else if (((p00 = (''fs1'')))) then True else if (((p00 = (''fa0'')))) then True else if (((p00 = (''fa1'')))) then True else if (((p00 = (''fa2'')))) then True else if (((p00 = (''fa3'')))) then True else if (((p00 = (''fa4'')))) then True else if (((p00 = (''fa5'')))) then True else if (((p00 = (''fa6'')))) then True else if (((p00 = (''fa7'')))) then True else if (((p00 = (''fs2'')))) then True else if (((p00 = (''fs3'')))) then True else if (((p00 = (''fs4'')))) then True else if (((p00 = (''fs5'')))) then True else if (((p00 = (''fs6'')))) then True else if (((p00 = (''fs7'')))) then True else if (((p00 = (''fs8'')))) then True else if (((p00 = (''fs9'')))) then True else if (((p00 = (''fs10'')))) then True else if (((p00 = (''fs11'')))) then True else if (((p00 = (''ft8'')))) then True else if (((p00 = (''ft9'')))) then True else if (((p00 = (''ft10'')))) then True else if (((p00 = (''ft11'')))) then True else False))\ for arg1 :: " string " \ \\val _s324_ : string -> maybe string\\ definition s324 :: \ string \(string)option \ where \ s324 s3250 = ( (let s3260 = s3250 in if ((string_startswith s3260 (''ft11''))) then (case ((string_drop s3260 ((string_length (''ft11''))))) of s1 => Some s1 ) else None))\ for s3250 :: " string " \ \\val _s320_ : string -> maybe string\\ definition s320 :: \ string \(string)option \ where \ s320 s3210 = ( (let s3220 = s3210 in if ((string_startswith s3220 (''ft10''))) then (case ((string_drop s3220 ((string_length (''ft10''))))) of s1 => Some s1 ) else None))\ for s3210 :: " string " \ \\val _s316_ : string -> maybe string\\ definition s316 :: \ string \(string)option \ where \ s316 s3170 = ( (let s3180 = s3170 in if ((string_startswith s3180 (''ft9''))) then (case ((string_drop s3180 ((string_length (''ft9''))))) of s1 => Some s1 ) else None))\ for s3170 :: " string " \ \\val _s312_ : string -> maybe string\\ definition s312 :: \ string \(string)option \ where \ s312 s3130 = ( (let s3140 = s3130 in if ((string_startswith s3140 (''ft8''))) then (case ((string_drop s3140 ((string_length (''ft8''))))) of s1 => Some s1 ) else None))\ for s3130 :: " string " \ \\val _s308_ : string -> maybe string\\ definition s308 :: \ string \(string)option \ where \ s308 s3090 = ( (let s3100 = s3090 in if ((string_startswith s3100 (''fs11''))) then (case ((string_drop s3100 ((string_length (''fs11''))))) of s1 => Some s1 ) else None))\ for s3090 :: " string " \ \\val _s304_ : string -> maybe string\\ definition s304 :: \ string \(string)option \ where \ s304 s3050 = ( (let s3060 = s3050 in if ((string_startswith s3060 (''fs10''))) then (case ((string_drop s3060 ((string_length (''fs10''))))) of s1 => Some s1 ) else None))\ for s3050 :: " string " \ \\val _s300_ : string -> maybe string\\ definition s300 :: \ string \(string)option \ where \ s300 s3010 = ( (let s3020 = s3010 in if ((string_startswith s3020 (''fs9''))) then (case ((string_drop s3020 ((string_length (''fs9''))))) of s1 => Some s1 ) else None))\ for s3010 :: " string " \ \\val _s296_ : string -> maybe string\\ definition s296 :: \ string \(string)option \ where \ s296 s2970 = ( (let s2980 = s2970 in if ((string_startswith s2980 (''fs8''))) then (case ((string_drop s2980 ((string_length (''fs8''))))) of s1 => Some s1 ) else None))\ for s2970 :: " string " \ \\val _s292_ : string -> maybe string\\ definition s292 :: \ string \(string)option \ where \ s292 s2930 = ( (let s2940 = s2930 in if ((string_startswith s2940 (''fs7''))) then (case ((string_drop s2940 ((string_length (''fs7''))))) of s1 => Some s1 ) else None))\ for s2930 :: " string " \ \\val _s288_ : string -> maybe string\\ definition s288 :: \ string \(string)option \ where \ s288 s2890 = ( (let s2900 = s2890 in if ((string_startswith s2900 (''fs6''))) then (case ((string_drop s2900 ((string_length (''fs6''))))) of s1 => Some s1 ) else None))\ for s2890 :: " string " \ \\val _s284_ : string -> maybe string\\ definition s284 :: \ string \(string)option \ where \ s284 s2850 = ( (let s2860 = s2850 in if ((string_startswith s2860 (''fs5''))) then (case ((string_drop s2860 ((string_length (''fs5''))))) of s1 => Some s1 ) else None))\ for s2850 :: " string " \ \\val _s280_ : string -> maybe string\\ definition s280 :: \ string \(string)option \ where \ s280 s2810 = ( (let s2821 = s2810 in if ((string_startswith s2821 (''fs4''))) then (case ((string_drop s2821 ((string_length (''fs4''))))) of s1 => Some s1 ) else None))\ for s2810 :: " string " \ \\val _s276_ : string -> maybe string\\ definition s276 :: \ string \(string)option \ where \ s276 s2770 = ( (let s2780 = s2770 in if ((string_startswith s2780 (''fs3''))) then (case ((string_drop s2780 ((string_length (''fs3''))))) of s1 => Some s1 ) else None))\ for s2770 :: " string " \ \\val _s272_ : string -> maybe string\\ definition s272 :: \ string \(string)option \ where \ s272 s2730 = ( (let s2740 = s2730 in if ((string_startswith s2740 (''fs2''))) then (case ((string_drop s2740 ((string_length (''fs2''))))) of s1 => Some s1 ) else None))\ for s2730 :: " string " \ \\val _s268_ : string -> maybe string\\ definition s268 :: \ string \(string)option \ where \ s268 s2690 = ( (let s2700 = s2690 in if ((string_startswith s2700 (''fa7''))) then (case ((string_drop s2700 ((string_length (''fa7''))))) of s1 => Some s1 ) else None))\ for s2690 :: " string " \ \\val _s264_ : string -> maybe string\\ definition s264 :: \ string \(string)option \ where \ s264 s2650 = ( (let s2660 = s2650 in if ((string_startswith s2660 (''fa6''))) then (case ((string_drop s2660 ((string_length (''fa6''))))) of s1 => Some s1 ) else None))\ for s2650 :: " string " \ \\val _s260_ : string -> maybe string\\ definition s260 :: \ string \(string)option \ where \ s260 s2610 = ( (let s2620 = s2610 in if ((string_startswith s2620 (''fa5''))) then (case ((string_drop s2620 ((string_length (''fa5''))))) of s1 => Some s1 ) else None))\ for s2610 :: " string " \ \\val _s256_ : string -> maybe string\\ definition s256 :: \ string \(string)option \ where \ s256 s2570 = ( (let s2580 = s2570 in if ((string_startswith s2580 (''fa4''))) then (case ((string_drop s2580 ((string_length (''fa4''))))) of s1 => Some s1 ) else None))\ for s2570 :: " string " \ \\val _s252_ : string -> maybe string\\ definition s252 :: \ string \(string)option \ where \ s252 s2530 = ( (let s2540 = s2530 in if ((string_startswith s2540 (''fa3''))) then (case ((string_drop s2540 ((string_length (''fa3''))))) of s1 => Some s1 ) else None))\ for s2530 :: " string " \ \\val _s248_ : string -> maybe string\\ definition s248 :: \ string \(string)option \ where \ s248 s2490 = ( (let s2500 = s2490 in if ((string_startswith s2500 (''fa2''))) then (case ((string_drop s2500 ((string_length (''fa2''))))) of s1 => Some s1 ) else None))\ for s2490 :: " string " \ \\val _s244_ : string -> maybe string\\ definition s244 :: \ string \(string)option \ where \ s244 s2450 = ( (let s2460 = s2450 in if ((string_startswith s2460 (''fa1''))) then (case ((string_drop s2460 ((string_length (''fa1''))))) of s1 => Some s1 ) else None))\ for s2450 :: " string " \ \\val _s240_ : string -> maybe string\\ definition s240 :: \ string \(string)option \ where \ s240 s2410 = ( (let s2420 = s2410 in if ((string_startswith s2420 (''fa0''))) then (case ((string_drop s2420 ((string_length (''fa0''))))) of s1 => Some s1 ) else None))\ for s2410 :: " string " \ \\val _s236_ : string -> maybe string\\ definition s236 :: \ string \(string)option \ where \ s236 s2370 = ( (let s2380 = s2370 in if ((string_startswith s2380 (''fs1''))) then (case ((string_drop s2380 ((string_length (''fs1''))))) of s1 => Some s1 ) else None))\ for s2370 :: " string " \ \\val _s232_ : string -> maybe string\\ definition s232 :: \ string \(string)option \ where \ s232 s2330 = ( (let s2340 = s2330 in if ((string_startswith s2340 (''fs0''))) then (case ((string_drop s2340 ((string_length (''fs0''))))) of s1 => Some s1 ) else None))\ for s2330 :: " string " \ \\val _s228_ : string -> maybe string\\ definition s228 :: \ string \(string)option \ where \ s228 s2290 = ( (let s2300 = s2290 in if ((string_startswith s2300 (''ft7''))) then (case ((string_drop s2300 ((string_length (''ft7''))))) of s1 => Some s1 ) else None))\ for s2290 :: " string " \ \\val _s224_ : string -> maybe string\\ definition s224 :: \ string \(string)option \ where \ s224 s2250 = ( (let s2260 = s2250 in if ((string_startswith s2260 (''ft6''))) then (case ((string_drop s2260 ((string_length (''ft6''))))) of s1 => Some s1 ) else None))\ for s2250 :: " string " \ \\val _s220_ : string -> maybe string\\ definition s220 :: \ string \(string)option \ where \ s220 s2210 = ( (let s2220 = s2210 in if ((string_startswith s2220 (''ft5''))) then (case ((string_drop s2220 ((string_length (''ft5''))))) of s1 => Some s1 ) else None))\ for s2210 :: " string " \ \\val _s216_ : string -> maybe string\\ definition s216 :: \ string \(string)option \ where \ s216 s2170 = ( (let s2180 = s2170 in if ((string_startswith s2180 (''ft4''))) then (case ((string_drop s2180 ((string_length (''ft4''))))) of s1 => Some s1 ) else None))\ for s2170 :: " string " \ \\val _s212_ : string -> maybe string\\ definition s212 :: \ string \(string)option \ where \ s212 s2130 = ( (let s2140 = s2130 in if ((string_startswith s2140 (''ft3''))) then (case ((string_drop s2140 ((string_length (''ft3''))))) of s1 => Some s1 ) else None))\ for s2130 :: " string " \ \\val _s208_ : string -> maybe string\\ definition s208 :: \ string \(string)option \ where \ s208 s2091 = ( (let s2100 = s2091 in if ((string_startswith s2100 (''ft2''))) then (case ((string_drop s2100 ((string_length (''ft2''))))) of s1 => Some s1 ) else None))\ for s2091 :: " string " \ \\val _s204_ : string -> maybe string\\ definition s204 :: \ string \(string)option \ where \ s204 s2050 = ( (let s2060 = s2050 in if ((string_startswith s2060 (''ft1''))) then (case ((string_drop s2060 ((string_length (''ft1''))))) of s1 => Some s1 ) else None))\ for s2050 :: " string " \ \\val _s200_ : string -> maybe string\\ definition s200 :: \ string \(string)option \ where \ s200 s2010 = ( (let s2020 = s2010 in if ((string_startswith s2020 (''ft0''))) then (case ((string_drop s2020 ((string_length (''ft0''))))) of s1 => Some s1 ) else None))\ for s2010 :: " string " definition freg_name_abi_matches_prefix :: \ string \((5)Word.word*int)option \ where \ freg_name_abi_matches_prefix arg1 = ( (let s2030 = arg1 in if ((case ((s200 s2030)) of Some (s1) => True | _ => False )) then (case s200 s2030 of (Some (s1)) => Some (( 0b00000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s204 s2030)) of Some (s1) => True | _ => False )) then (case s204 s2030 of (Some (s1)) => Some (( 0b00001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s208 s2030)) of Some (s1) => True | _ => False )) then (case s208 s2030 of (Some (s1)) => Some (( 0b00010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s212 s2030)) of Some (s1) => True | _ => False )) then (case s212 s2030 of (Some (s1)) => Some (( 0b00011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s216 s2030)) of Some (s1) => True | _ => False )) then (case s216 s2030 of (Some (s1)) => Some (( 0b00100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s220 s2030)) of Some (s1) => True | _ => False )) then (case s220 s2030 of (Some (s1)) => Some (( 0b00101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s224 s2030)) of Some (s1) => True | _ => False )) then (case s224 s2030 of (Some (s1)) => Some (( 0b00110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s228 s2030)) of Some (s1) => True | _ => False )) then (case s228 s2030 of (Some (s1)) => Some (( 0b00111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s232 s2030)) of Some (s1) => True | _ => False )) then (case s232 s2030 of (Some (s1)) => Some (( 0b01000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s236 s2030)) of Some (s1) => True | _ => False )) then (case s236 s2030 of (Some (s1)) => Some (( 0b01001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s240 s2030)) of Some (s1) => True | _ => False )) then (case s240 s2030 of (Some (s1)) => Some (( 0b01010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s244 s2030)) of Some (s1) => True | _ => False )) then (case s244 s2030 of (Some (s1)) => Some (( 0b01011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s248 s2030)) of Some (s1) => True | _ => False )) then (case s248 s2030 of (Some (s1)) => Some (( 0b01100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s252 s2030)) of Some (s1) => True | _ => False )) then (case s252 s2030 of (Some (s1)) => Some (( 0b01101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s256 s2030)) of Some (s1) => True | _ => False )) then (case s256 s2030 of (Some (s1)) => Some (( 0b01110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s260 s2030)) of Some (s1) => True | _ => False )) then (case s260 s2030 of (Some (s1)) => Some (( 0b01111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s264 s2030)) of Some (s1) => True | _ => False )) then (case s264 s2030 of (Some (s1)) => Some (( 0b10000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s268 s2030)) of Some (s1) => True | _ => False )) then (case s268 s2030 of (Some (s1)) => Some (( 0b10001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s272 s2030)) of Some (s1) => True | _ => False )) then (case s272 s2030 of (Some (s1)) => Some (( 0b10010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s276 s2030)) of Some (s1) => True | _ => False )) then (case s276 s2030 of (Some (s1)) => Some (( 0b10011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s280 s2030)) of Some (s1) => True | _ => False )) then (case s280 s2030 of (Some (s1)) => Some (( 0b10100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s284 s2030)) of Some (s1) => True | _ => False )) then (case s284 s2030 of (Some (s1)) => Some (( 0b10101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s288 s2030)) of Some (s1) => True | _ => False )) then (case s288 s2030 of (Some (s1)) => Some (( 0b10110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s292 s2030)) of Some (s1) => True | _ => False )) then (case s292 s2030 of (Some (s1)) => Some (( 0b10111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s296 s2030)) of Some (s1) => True | _ => False )) then (case s296 s2030 of (Some (s1)) => Some (( 0b11000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s300 s2030)) of Some (s1) => True | _ => False )) then (case s300 s2030 of (Some (s1)) => Some (( 0b11001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s304 s2030)) of Some (s1) => True | _ => False )) then (case s304 s2030 of (Some (s1)) => Some (( 0b11010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s308 s2030)) of Some (s1) => True | _ => False )) then (case s308 s2030 of (Some (s1)) => Some (( 0b11011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s312 s2030)) of Some (s1) => True | _ => False )) then (case s312 s2030 of (Some (s1)) => Some (( 0b11100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s316 s2030)) of Some (s1) => True | _ => False )) then (case s316 s2030 of (Some (s1)) => Some (( 0b11101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s320 s2030)) of Some (s1) => True | _ => False )) then (case s320 s2030 of (Some (s1)) => Some (( 0b11110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s324 s2030)) of Some (s1) => True | _ => False )) then (case s324 s2030 of (Some (s1)) => Some (( 0b11111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val freg_name_forwards : mword ty5 -> M string\\ \ \\val freg_name_backwards : string -> M (mword ty5)\\ \ \\val freg_name_forwards_matches : mword ty5 -> bool\\ \ \\val freg_name_backwards_matches : string -> bool\\ \ \\val freg_name_matches_prefix : string -> maybe ((mword ty5 * ii))\\ definition freg_name_forwards :: \(5)Word.word \((register_value),(string),(exception))monad \ where \ freg_name_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00000 :: 5 Word.word)))) then return (''ft0'') else if (((b__0 = ( 0b00001 :: 5 Word.word)))) then return (''ft1'') else if (((b__0 = ( 0b00010 :: 5 Word.word)))) then return (''ft2'') else if (((b__0 = ( 0b00011 :: 5 Word.word)))) then return (''ft3'') else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then return (''ft4'') else if (((b__0 = ( 0b00101 :: 5 Word.word)))) then return (''ft5'') else if (((b__0 = ( 0b00110 :: 5 Word.word)))) then return (''ft6'') else if (((b__0 = ( 0b00111 :: 5 Word.word)))) then return (''ft7'') else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then return (''fs0'') else if (((b__0 = ( 0b01001 :: 5 Word.word)))) then return (''fs1'') else if (((b__0 = ( 0b01010 :: 5 Word.word)))) then return (''fa0'') else if (((b__0 = ( 0b01011 :: 5 Word.word)))) then return (''fa1'') else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then return (''fa2'') else if (((b__0 = ( 0b01101 :: 5 Word.word)))) then return (''fa3'') else if (((b__0 = ( 0b01110 :: 5 Word.word)))) then return (''fa4'') else if (((b__0 = ( 0b01111 :: 5 Word.word)))) then return (''fa5'') else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then return (''fa6'') else if (((b__0 = ( 0b10001 :: 5 Word.word)))) then return (''fa7'') else if (((b__0 = ( 0b10010 :: 5 Word.word)))) then return (''fs2'') else if (((b__0 = ( 0b10011 :: 5 Word.word)))) then return (''fs3'') else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then return (''fs4'') else if (((b__0 = ( 0b10101 :: 5 Word.word)))) then return (''fs5'') else if (((b__0 = ( 0b10110 :: 5 Word.word)))) then return (''fs6'') else if (((b__0 = ( 0b10111 :: 5 Word.word)))) then return (''fs7'') else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then return (''fs8'') else if (((b__0 = ( 0b11001 :: 5 Word.word)))) then return (''fs9'') else if (((b__0 = ( 0b11010 :: 5 Word.word)))) then return (''fs10'') else if (((b__0 = ( 0b11011 :: 5 Word.word)))) then return (''fs11'') else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then return (''ft8'') else if (((b__0 = ( 0b11101 :: 5 Word.word)))) then return (''ft9'') else if (((b__0 = ( 0b11110 :: 5 Word.word)))) then return (''ft10'') else if (((b__0 = ( 0b11111 :: 5 Word.word)))) then return (''ft11'') else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(5)Word.word " definition freg_name_backwards :: \ string \((register_value),((5)Word.word),(exception))monad \ where \ freg_name_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''ft0'')))) then return ( 0b00000 :: 5 Word.word) else if (((p00 = (''ft1'')))) then return ( 0b00001 :: 5 Word.word) else if (((p00 = (''ft2'')))) then return ( 0b00010 :: 5 Word.word) else if (((p00 = (''ft3'')))) then return ( 0b00011 :: 5 Word.word) else if (((p00 = (''ft4'')))) then return ( 0b00100 :: 5 Word.word) else if (((p00 = (''ft5'')))) then return ( 0b00101 :: 5 Word.word) else if (((p00 = (''ft6'')))) then return ( 0b00110 :: 5 Word.word) else if (((p00 = (''ft7'')))) then return ( 0b00111 :: 5 Word.word) else if (((p00 = (''fs0'')))) then return ( 0b01000 :: 5 Word.word) else if (((p00 = (''fs1'')))) then return ( 0b01001 :: 5 Word.word) else if (((p00 = (''fa0'')))) then return ( 0b01010 :: 5 Word.word) else if (((p00 = (''fa1'')))) then return ( 0b01011 :: 5 Word.word) else if (((p00 = (''fa2'')))) then return ( 0b01100 :: 5 Word.word) else if (((p00 = (''fa3'')))) then return ( 0b01101 :: 5 Word.word) else if (((p00 = (''fa4'')))) then return ( 0b01110 :: 5 Word.word) else if (((p00 = (''fa5'')))) then return ( 0b01111 :: 5 Word.word) else if (((p00 = (''fa6'')))) then return ( 0b10000 :: 5 Word.word) else if (((p00 = (''fa7'')))) then return ( 0b10001 :: 5 Word.word) else if (((p00 = (''fs2'')))) then return ( 0b10010 :: 5 Word.word) else if (((p00 = (''fs3'')))) then return ( 0b10011 :: 5 Word.word) else if (((p00 = (''fs4'')))) then return ( 0b10100 :: 5 Word.word) else if (((p00 = (''fs5'')))) then return ( 0b10101 :: 5 Word.word) else if (((p00 = (''fs6'')))) then return ( 0b10110 :: 5 Word.word) else if (((p00 = (''fs7'')))) then return ( 0b10111 :: 5 Word.word) else if (((p00 = (''fs8'')))) then return ( 0b11000 :: 5 Word.word) else if (((p00 = (''fs9'')))) then return ( 0b11001 :: 5 Word.word) else if (((p00 = (''fs10'')))) then return ( 0b11010 :: 5 Word.word) else if (((p00 = (''fs11'')))) then return ( 0b11011 :: 5 Word.word) else if (((p00 = (''ft8'')))) then return ( 0b11100 :: 5 Word.word) else if (((p00 = (''ft9'')))) then return ( 0b11101 :: 5 Word.word) else if (((p00 = (''ft10'')))) then return ( 0b11110 :: 5 Word.word) else if (((p00 = (''ft11'')))) then return ( 0b11111 :: 5 Word.word) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition freg_name_forwards_matches :: \(5)Word.word \ bool \ where \ freg_name_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10111 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11010 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11011 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11101 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11110 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11111 :: 5 Word.word)))) then True else False))\ for arg1 :: "(5)Word.word " definition freg_name_backwards_matches :: \ string \ bool \ where \ freg_name_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''ft0'')))) then True else if (((p00 = (''ft1'')))) then True else if (((p00 = (''ft2'')))) then True else if (((p00 = (''ft3'')))) then True else if (((p00 = (''ft4'')))) then True else if (((p00 = (''ft5'')))) then True else if (((p00 = (''ft6'')))) then True else if (((p00 = (''ft7'')))) then True else if (((p00 = (''fs0'')))) then True else if (((p00 = (''fs1'')))) then True else if (((p00 = (''fa0'')))) then True else if (((p00 = (''fa1'')))) then True else if (((p00 = (''fa2'')))) then True else if (((p00 = (''fa3'')))) then True else if (((p00 = (''fa4'')))) then True else if (((p00 = (''fa5'')))) then True else if (((p00 = (''fa6'')))) then True else if (((p00 = (''fa7'')))) then True else if (((p00 = (''fs2'')))) then True else if (((p00 = (''fs3'')))) then True else if (((p00 = (''fs4'')))) then True else if (((p00 = (''fs5'')))) then True else if (((p00 = (''fs6'')))) then True else if (((p00 = (''fs7'')))) then True else if (((p00 = (''fs8'')))) then True else if (((p00 = (''fs9'')))) then True else if (((p00 = (''fs10'')))) then True else if (((p00 = (''fs11'')))) then True else if (((p00 = (''ft8'')))) then True else if (((p00 = (''ft9'')))) then True else if (((p00 = (''ft10'')))) then True else if (((p00 = (''ft11'')))) then True else False))\ for arg1 :: " string " \ \\val _s452_ : string -> maybe string\\ definition s452 :: \ string \(string)option \ where \ s452 s4530 = ( (let s4540 = s4530 in if ((string_startswith s4540 (''ft11''))) then (case ((string_drop s4540 ((string_length (''ft11''))))) of s1 => Some s1 ) else None))\ for s4530 :: " string " \ \\val _s448_ : string -> maybe string\\ definition s448 :: \ string \(string)option \ where \ s448 s4490 = ( (let s4501 = s4490 in if ((string_startswith s4501 (''ft10''))) then (case ((string_drop s4501 ((string_length (''ft10''))))) of s1 => Some s1 ) else None))\ for s4490 :: " string " \ \\val _s444_ : string -> maybe string\\ definition s444 :: \ string \(string)option \ where \ s444 s4450 = ( (let s4460 = s4450 in if ((string_startswith s4460 (''ft9''))) then (case ((string_drop s4460 ((string_length (''ft9''))))) of s1 => Some s1 ) else None))\ for s4450 :: " string " \ \\val _s440_ : string -> maybe string\\ definition s440 :: \ string \(string)option \ where \ s440 s4410 = ( (let s4420 = s4410 in if ((string_startswith s4420 (''ft8''))) then (case ((string_drop s4420 ((string_length (''ft8''))))) of s1 => Some s1 ) else None))\ for s4410 :: " string " \ \\val _s436_ : string -> maybe string\\ definition s436 :: \ string \(string)option \ where \ s436 s4371 = ( (let s4380 = s4371 in if ((string_startswith s4380 (''fs11''))) then (case ((string_drop s4380 ((string_length (''fs11''))))) of s1 => Some s1 ) else None))\ for s4371 :: " string " \ \\val _s432_ : string -> maybe string\\ definition s432 :: \ string \(string)option \ where \ s432 s4330 = ( (let s4340 = s4330 in if ((string_startswith s4340 (''fs10''))) then (case ((string_drop s4340 ((string_length (''fs10''))))) of s1 => Some s1 ) else None))\ for s4330 :: " string " \ \\val _s428_ : string -> maybe string\\ definition s428 :: \ string \(string)option \ where \ s428 s4290 = ( (let s4300 = s4290 in if ((string_startswith s4300 (''fs9''))) then (case ((string_drop s4300 ((string_length (''fs9''))))) of s1 => Some s1 ) else None))\ for s4290 :: " string " \ \\val _s424_ : string -> maybe string\\ definition s424 :: \ string \(string)option \ where \ s424 s4250 = ( (let s4260 = s4250 in if ((string_startswith s4260 (''fs8''))) then (case ((string_drop s4260 ((string_length (''fs8''))))) of s1 => Some s1 ) else None))\ for s4250 :: " string " \ \\val _s420_ : string -> maybe string\\ definition s420 :: \ string \(string)option \ where \ s420 s4210 = ( (let s4220 = s4210 in if ((string_startswith s4220 (''fs7''))) then (case ((string_drop s4220 ((string_length (''fs7''))))) of s1 => Some s1 ) else None))\ for s4210 :: " string " \ \\val _s416_ : string -> maybe string\\ definition s416 :: \ string \(string)option \ where \ s416 s4170 = ( (let s4180 = s4170 in if ((string_startswith s4180 (''fs6''))) then (case ((string_drop s4180 ((string_length (''fs6''))))) of s1 => Some s1 ) else None))\ for s4170 :: " string " \ \\val _s412_ : string -> maybe string\\ definition s412 :: \ string \(string)option \ where \ s412 s4130 = ( (let s4140 = s4130 in if ((string_startswith s4140 (''fs5''))) then (case ((string_drop s4140 ((string_length (''fs5''))))) of s1 => Some s1 ) else None))\ for s4130 :: " string " \ \\val _s408_ : string -> maybe string\\ definition s408 :: \ string \(string)option \ where \ s408 s4090 = ( (let s4100 = s4090 in if ((string_startswith s4100 (''fs4''))) then (case ((string_drop s4100 ((string_length (''fs4''))))) of s1 => Some s1 ) else None))\ for s4090 :: " string " \ \\val _s404_ : string -> maybe string\\ definition s404 :: \ string \(string)option \ where \ s404 s4050 = ( (let s4061 = s4050 in if ((string_startswith s4061 (''fs3''))) then (case ((string_drop s4061 ((string_length (''fs3''))))) of s1 => Some s1 ) else None))\ for s4050 :: " string " \ \\val _s400_ : string -> maybe string\\ definition s400 :: \ string \(string)option \ where \ s400 s4010 = ( (let s4021 = s4010 in if ((string_startswith s4021 (''fs2''))) then (case ((string_drop s4021 ((string_length (''fs2''))))) of s1 => Some s1 ) else None))\ for s4010 :: " string " \ \\val _s396_ : string -> maybe string\\ definition s396 :: \ string \(string)option \ where \ s396 s3970 = ( (let s3980 = s3970 in if ((string_startswith s3980 (''fa7''))) then (case ((string_drop s3980 ((string_length (''fa7''))))) of s1 => Some s1 ) else None))\ for s3970 :: " string " \ \\val _s392_ : string -> maybe string\\ definition s392 :: \ string \(string)option \ where \ s392 s3931 = ( (let s3940 = s3931 in if ((string_startswith s3940 (''fa6''))) then (case ((string_drop s3940 ((string_length (''fa6''))))) of s1 => Some s1 ) else None))\ for s3931 :: " string " \ \\val _s388_ : string -> maybe string\\ definition s388 :: \ string \(string)option \ where \ s388 s3890 = ( (let s3900 = s3890 in if ((string_startswith s3900 (''fa5''))) then (case ((string_drop s3900 ((string_length (''fa5''))))) of s1 => Some s1 ) else None))\ for s3890 :: " string " \ \\val _s384_ : string -> maybe string\\ definition s384 :: \ string \(string)option \ where \ s384 s3850 = ( (let s3860 = s3850 in if ((string_startswith s3860 (''fa4''))) then (case ((string_drop s3860 ((string_length (''fa4''))))) of s1 => Some s1 ) else None))\ for s3850 :: " string " \ \\val _s380_ : string -> maybe string\\ definition s380 :: \ string \(string)option \ where \ s380 s3810 = ( (let s3820 = s3810 in if ((string_startswith s3820 (''fa3''))) then (case ((string_drop s3820 ((string_length (''fa3''))))) of s1 => Some s1 ) else None))\ for s3810 :: " string " \ \\val _s376_ : string -> maybe string\\ definition s376 :: \ string \(string)option \ where \ s376 s3770 = ( (let s3780 = s3770 in if ((string_startswith s3780 (''fa2''))) then (case ((string_drop s3780 ((string_length (''fa2''))))) of s1 => Some s1 ) else None))\ for s3770 :: " string " \ \\val _s372_ : string -> maybe string\\ definition s372 :: \ string \(string)option \ where \ s372 s3730 = ( (let s3740 = s3730 in if ((string_startswith s3740 (''fa1''))) then (case ((string_drop s3740 ((string_length (''fa1''))))) of s1 => Some s1 ) else None))\ for s3730 :: " string " \ \\val _s368_ : string -> maybe string\\ definition s368 :: \ string \(string)option \ where \ s368 s3690 = ( (let s3700 = s3690 in if ((string_startswith s3700 (''fa0''))) then (case ((string_drop s3700 ((string_length (''fa0''))))) of s1 => Some s1 ) else None))\ for s3690 :: " string " \ \\val _s364_ : string -> maybe string\\ definition s364 :: \ string \(string)option \ where \ s364 s3650 = ( (let s3660 = s3650 in if ((string_startswith s3660 (''fs1''))) then (case ((string_drop s3660 ((string_length (''fs1''))))) of s1 => Some s1 ) else None))\ for s3650 :: " string " \ \\val _s360_ : string -> maybe string\\ definition s360 :: \ string \(string)option \ where \ s360 s3610 = ( (let s3620 = s3610 in if ((string_startswith s3620 (''fs0''))) then (case ((string_drop s3620 ((string_length (''fs0''))))) of s1 => Some s1 ) else None))\ for s3610 :: " string " \ \\val _s356_ : string -> maybe string\\ definition s356 :: \ string \(string)option \ where \ s356 s3570 = ( (let s3580 = s3570 in if ((string_startswith s3580 (''ft7''))) then (case ((string_drop s3580 ((string_length (''ft7''))))) of s1 => Some s1 ) else None))\ for s3570 :: " string " \ \\val _s352_ : string -> maybe string\\ definition s352 :: \ string \(string)option \ where \ s352 s3530 = ( (let s3540 = s3530 in if ((string_startswith s3540 (''ft6''))) then (case ((string_drop s3540 ((string_length (''ft6''))))) of s1 => Some s1 ) else None))\ for s3530 :: " string " \ \\val _s348_ : string -> maybe string\\ definition s348 :: \ string \(string)option \ where \ s348 s3490 = ( (let s3500 = s3490 in if ((string_startswith s3500 (''ft5''))) then (case ((string_drop s3500 ((string_length (''ft5''))))) of s1 => Some s1 ) else None))\ for s3490 :: " string " \ \\val _s344_ : string -> maybe string\\ definition s344 :: \ string \(string)option \ where \ s344 s3450 = ( (let s3460 = s3450 in if ((string_startswith s3460 (''ft4''))) then (case ((string_drop s3460 ((string_length (''ft4''))))) of s1 => Some s1 ) else None))\ for s3450 :: " string " \ \\val _s340_ : string -> maybe string\\ definition s340 :: \ string \(string)option \ where \ s340 s3410 = ( (let s3420 = s3410 in if ((string_startswith s3420 (''ft3''))) then (case ((string_drop s3420 ((string_length (''ft3''))))) of s1 => Some s1 ) else None))\ for s3410 :: " string " \ \\val _s336_ : string -> maybe string\\ definition s336 :: \ string \(string)option \ where \ s336 s3370 = ( (let s3380 = s3370 in if ((string_startswith s3380 (''ft2''))) then (case ((string_drop s3380 ((string_length (''ft2''))))) of s1 => Some s1 ) else None))\ for s3370 :: " string " \ \\val _s332_ : string -> maybe string\\ definition s332 :: \ string \(string)option \ where \ s332 s3330 = ( (let s3340 = s3330 in if ((string_startswith s3340 (''ft1''))) then (case ((string_drop s3340 ((string_length (''ft1''))))) of s1 => Some s1 ) else None))\ for s3330 :: " string " \ \\val _s328_ : string -> maybe string\\ definition s328 :: \ string \(string)option \ where \ s328 s3290 = ( (let s3300 = s3290 in if ((string_startswith s3300 (''ft0''))) then (case ((string_drop s3300 ((string_length (''ft0''))))) of s1 => Some s1 ) else None))\ for s3290 :: " string " definition freg_name_matches_prefix :: \ string \((5)Word.word*int)option \ where \ freg_name_matches_prefix arg1 = ( (let s3310 = arg1 in if ((case ((s328 s3310)) of Some (s1) => True | _ => False )) then (case s328 s3310 of (Some (s1)) => Some (( 0b00000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s332 s3310)) of Some (s1) => True | _ => False )) then (case s332 s3310 of (Some (s1)) => Some (( 0b00001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s336 s3310)) of Some (s1) => True | _ => False )) then (case s336 s3310 of (Some (s1)) => Some (( 0b00010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s340 s3310)) of Some (s1) => True | _ => False )) then (case s340 s3310 of (Some (s1)) => Some (( 0b00011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s344 s3310)) of Some (s1) => True | _ => False )) then (case s344 s3310 of (Some (s1)) => Some (( 0b00100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s348 s3310)) of Some (s1) => True | _ => False )) then (case s348 s3310 of (Some (s1)) => Some (( 0b00101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s352 s3310)) of Some (s1) => True | _ => False )) then (case s352 s3310 of (Some (s1)) => Some (( 0b00110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s356 s3310)) of Some (s1) => True | _ => False )) then (case s356 s3310 of (Some (s1)) => Some (( 0b00111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s360 s3310)) of Some (s1) => True | _ => False )) then (case s360 s3310 of (Some (s1)) => Some (( 0b01000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s364 s3310)) of Some (s1) => True | _ => False )) then (case s364 s3310 of (Some (s1)) => Some (( 0b01001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s368 s3310)) of Some (s1) => True | _ => False )) then (case s368 s3310 of (Some (s1)) => Some (( 0b01010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s372 s3310)) of Some (s1) => True | _ => False )) then (case s372 s3310 of (Some (s1)) => Some (( 0b01011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s376 s3310)) of Some (s1) => True | _ => False )) then (case s376 s3310 of (Some (s1)) => Some (( 0b01100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s380 s3310)) of Some (s1) => True | _ => False )) then (case s380 s3310 of (Some (s1)) => Some (( 0b01101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s384 s3310)) of Some (s1) => True | _ => False )) then (case s384 s3310 of (Some (s1)) => Some (( 0b01110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s388 s3310)) of Some (s1) => True | _ => False )) then (case s388 s3310 of (Some (s1)) => Some (( 0b01111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s392 s3310)) of Some (s1) => True | _ => False )) then (case s392 s3310 of (Some (s1)) => Some (( 0b10000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s396 s3310)) of Some (s1) => True | _ => False )) then (case s396 s3310 of (Some (s1)) => Some (( 0b10001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s400 s3310)) of Some (s1) => True | _ => False )) then (case s400 s3310 of (Some (s1)) => Some (( 0b10010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s404 s3310)) of Some (s1) => True | _ => False )) then (case s404 s3310 of (Some (s1)) => Some (( 0b10011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s408 s3310)) of Some (s1) => True | _ => False )) then (case s408 s3310 of (Some (s1)) => Some (( 0b10100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s412 s3310)) of Some (s1) => True | _ => False )) then (case s412 s3310 of (Some (s1)) => Some (( 0b10101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s416 s3310)) of Some (s1) => True | _ => False )) then (case s416 s3310 of (Some (s1)) => Some (( 0b10110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s420 s3310)) of Some (s1) => True | _ => False )) then (case s420 s3310 of (Some (s1)) => Some (( 0b10111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s424 s3310)) of Some (s1) => True | _ => False )) then (case s424 s3310 of (Some (s1)) => Some (( 0b11000 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s428 s3310)) of Some (s1) => True | _ => False )) then (case s428 s3310 of (Some (s1)) => Some (( 0b11001 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s432 s3310)) of Some (s1) => True | _ => False )) then (case s432 s3310 of (Some (s1)) => Some (( 0b11010 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s436 s3310)) of Some (s1) => True | _ => False )) then (case s436 s3310 of (Some (s1)) => Some (( 0b11011 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s440 s3310)) of Some (s1) => True | _ => False )) then (case s440 s3310 of (Some (s1)) => Some (( 0b11100 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s444 s3310)) of Some (s1) => True | _ => False )) then (case s444 s3310 of (Some (s1)) => Some (( 0b11101 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s448 s3310)) of Some (s1) => True | _ => False )) then (case s448 s3310 of (Some (s1)) => Some (( 0b11110 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s452 s3310)) of Some (s1) => True | _ => False )) then (case s452 s3310 of (Some (s1)) => Some (( 0b11111 :: 5 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val init_fdext_regs : unit -> M unit\\ definition init_fdext_regs :: \ unit \((register_value),(unit),(exception))monad \ where \ init_fdext_regs _ = ( ((((((((((((((((((((((((((((((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 )\ \ \\val undefined_Fcsr : unit -> M Fcsr\\ definition undefined_Fcsr :: \ unit \((register_value),(Fcsr),(exception))monad \ where \ undefined_Fcsr _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| Fcsr_bits = w__0 |)))))\ \ \\val Mk_Fcsr : mword ty32 -> Fcsr\\ definition Mk_Fcsr :: \(32)Word.word \ Fcsr \ where \ Mk_Fcsr v = ( (| Fcsr_bits = v |) )\ for v :: "(32)Word.word " definition get_Fcsr_bits :: \ Fcsr \(32)Word.word \ where \ get_Fcsr_bits v = ( (subrange_vec_dec(Fcsr_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " Fcsr " definition set_Fcsr_bits :: \((regstate),(register_value),(Fcsr))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_Fcsr_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Fcsr_bits := ((update_subrange_vec_dec(Fcsr_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Fcsr))register_ref " and v :: "(32)Word.word " definition update_Fcsr_bits :: \ Fcsr \(32)Word.word \ Fcsr \ where \ update_Fcsr_bits v x = ( ( v (| Fcsr_bits := ((update_subrange_vec_dec(Fcsr_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Fcsr " and x :: "(32)Word.word " \ \\val _get_Fcsr_FFLAGS : Fcsr -> mword ty5\\ definition get_Fcsr_FFLAGS :: \ Fcsr \(5)Word.word \ where \ get_Fcsr_FFLAGS v = ( (subrange_vec_dec(Fcsr_bits v) (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word))\ for v :: " Fcsr " \ \\val _set_Fcsr_FFLAGS : register_ref regstate register_value Fcsr -> mword ty5 -> M unit\\ definition set_Fcsr_FFLAGS :: \((regstate),(register_value),(Fcsr))register_ref \(5)Word.word \((register_value),(unit),(exception))monad \ where \ set_Fcsr_FFLAGS r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Fcsr_bits := ((update_subrange_vec_dec(Fcsr_bits r) (( 4 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Fcsr))register_ref " and v :: "(5)Word.word " \ \\val _update_Fcsr_FFLAGS : Fcsr -> mword ty5 -> Fcsr\\ definition update_Fcsr_FFLAGS :: \ Fcsr \(5)Word.word \ Fcsr \ where \ update_Fcsr_FFLAGS v x = ( ( v (| Fcsr_bits := ((update_subrange_vec_dec(Fcsr_bits v) (( 4 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Fcsr " and x :: "(5)Word.word " \ \\val _get_Fcsr_FRM : Fcsr -> mword ty3\\ definition get_Fcsr_FRM :: \ Fcsr \(3)Word.word \ where \ get_Fcsr_FRM v = ( (subrange_vec_dec(Fcsr_bits v) (( 7 :: int)::ii) (( 5 :: int)::ii) :: 3 Word.word))\ for v :: " Fcsr " \ \\val _set_Fcsr_FRM : register_ref regstate register_value Fcsr -> mword ty3 -> M unit\\ definition set_Fcsr_FRM :: \((regstate),(register_value),(Fcsr))register_ref \(3)Word.word \((register_value),(unit),(exception))monad \ where \ set_Fcsr_FRM r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| Fcsr_bits := ((update_subrange_vec_dec(Fcsr_bits r) (( 7 :: int)::ii) (( 5 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(Fcsr))register_ref " and v :: "(3)Word.word " \ \\val _update_Fcsr_FRM : Fcsr -> mword ty3 -> Fcsr\\ definition update_Fcsr_FRM :: \ Fcsr \(3)Word.word \ Fcsr \ where \ update_Fcsr_FRM v x = ( ( v (| Fcsr_bits := ((update_subrange_vec_dec(Fcsr_bits v) (( 7 :: int)::ii) (( 5 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " Fcsr " and x :: "(3)Word.word " \ \\val ext_write_fcsr : mword ty3 -> mword ty5 -> M unit\\ definition ext_write_fcsr :: \(3)Word.word \(5)Word.word \((register_value),(unit),(exception))monad \ where \ ext_write_fcsr frm fflags = ( ((set_Fcsr_FRM fcsr_ref frm \ set_Fcsr_FFLAGS fcsr_ref fflags) \ update_softfloat_fflags fflags) \ dirty_fd_context () )\ for frm :: "(3)Word.word " and fflags :: "(5)Word.word " \ \\val write_fflags : mword ty5 -> M unit\\ definition write_fflags :: \(5)Word.word \((register_value),(unit),(exception))monad \ where \ write_fflags fflags = ( read_reg fcsr_ref \ ((\ (w__0 :: Fcsr) . (if (((((get_Fcsr_FFLAGS w__0 :: 5 Word.word)) \ fflags))) then dirty_fd_context () else return () ) \ set_Fcsr_FFLAGS fcsr_ref fflags)))\ for fflags :: "(5)Word.word " \ \\val accrue_fflags : mword ty5 -> M unit\\ definition accrue_fflags :: \(5)Word.word \((register_value),(unit),(exception))monad \ where \ accrue_fflags flags = ( read_reg fcsr_ref \ ((\ (w__0 :: Fcsr) . (let f = ((or_vec ((get_Fcsr_FFLAGS w__0 :: 5 Word.word)) flags :: 5 Word.word)) in read_reg fcsr_ref \ ((\ (w__1 :: Fcsr) . if (((((get_Fcsr_FFLAGS w__1 :: 5 Word.word)) \ f))) then (set_Fcsr_FFLAGS fcsr_ref f \ update_softfloat_fflags f) \ dirty_fd_context () else return () ))))))\ for flags :: "(5)Word.word " definition csr_name_map_forwards :: \(12)Word.word \ string \ where \ csr_name_map_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0x000 :: 12 Word.word)))) then (''ustatus'') else if (((b__0 = ( 0x004 :: 12 Word.word)))) then (''uie'') else if (((b__0 = ( 0x005 :: 12 Word.word)))) then (''utvec'') else if (((b__0 = ( 0x040 :: 12 Word.word)))) then (''uscratch'') else if (((b__0 = ( 0x041 :: 12 Word.word)))) then (''uepc'') else if (((b__0 = ( 0x042 :: 12 Word.word)))) then (''ucause'') else if (((b__0 = ( 0x043 :: 12 Word.word)))) then (''utval'') else if (((b__0 = ( 0x044 :: 12 Word.word)))) then (''uip'') else if (((b__0 = ( 0x001 :: 12 Word.word)))) then (''fflags'') else if (((b__0 = ( 0x002 :: 12 Word.word)))) then (''frm'') else if (((b__0 = ( 0x003 :: 12 Word.word)))) then (''fcsr'') else if (((b__0 = ( 0xC00 :: 12 Word.word)))) then (''cycle'') else if (((b__0 = ( 0xC01 :: 12 Word.word)))) then (''time'') else if (((b__0 = ( 0xC02 :: 12 Word.word)))) then (''instret'') else if (((b__0 = ( 0xC80 :: 12 Word.word)))) then (''cycleh'') else if (((b__0 = ( 0xC81 :: 12 Word.word)))) then (''timeh'') else if (((b__0 = ( 0xC82 :: 12 Word.word)))) then (''instreth'') else if (((b__0 = ( 0x100 :: 12 Word.word)))) then (''sstatus'') else if (((b__0 = ( 0x102 :: 12 Word.word)))) then (''sedeleg'') else if (((b__0 = ( 0x103 :: 12 Word.word)))) then (''sideleg'') else if (((b__0 = ( 0x104 :: 12 Word.word)))) then (''sie'') else if (((b__0 = ( 0x105 :: 12 Word.word)))) then (''stvec'') else if (((b__0 = ( 0x106 :: 12 Word.word)))) then (''scounteren'') else if (((b__0 = ( 0x140 :: 12 Word.word)))) then (''sscratch'') else if (((b__0 = ( 0x141 :: 12 Word.word)))) then (''sepc'') else if (((b__0 = ( 0x142 :: 12 Word.word)))) then (''scause'') else if (((b__0 = ( 0x143 :: 12 Word.word)))) then (''stval'') else if (((b__0 = ( 0x144 :: 12 Word.word)))) then (''sip'') else if (((b__0 = ( 0x180 :: 12 Word.word)))) then (''satp'') else if (((b__0 = ( 0xF11 :: 12 Word.word)))) then (''mvendorid'') else if (((b__0 = ( 0xF12 :: 12 Word.word)))) then (''marchid'') else if (((b__0 = ( 0xF13 :: 12 Word.word)))) then (''mimpid'') else if (((b__0 = ( 0xF14 :: 12 Word.word)))) then (''mhartid'') else if (((b__0 = ( 0x300 :: 12 Word.word)))) then (''mstatus'') else if (((b__0 = ( 0x301 :: 12 Word.word)))) then (''misa'') else if (((b__0 = ( 0x302 :: 12 Word.word)))) then (''medeleg'') else if (((b__0 = ( 0x303 :: 12 Word.word)))) then (''mideleg'') else if (((b__0 = ( 0x304 :: 12 Word.word)))) then (''mie'') else if (((b__0 = ( 0x305 :: 12 Word.word)))) then (''mtvec'') else if (((b__0 = ( 0x306 :: 12 Word.word)))) then (''mcounteren'') else if (((b__0 = ( 0x320 :: 12 Word.word)))) then (''mcountinhibit'') else if (((b__0 = ( 0x340 :: 12 Word.word)))) then (''mscratch'') else if (((b__0 = ( 0x341 :: 12 Word.word)))) then (''mepc'') else if (((b__0 = ( 0x342 :: 12 Word.word)))) then (''mcause'') else if (((b__0 = ( 0x343 :: 12 Word.word)))) then (''mtval'') else if (((b__0 = ( 0x344 :: 12 Word.word)))) then (''mip'') else if (((b__0 = ( 0x3A0 :: 12 Word.word)))) then (''pmpcfg0'') else if (((b__0 = ( 0x3A1 :: 12 Word.word)))) then (''pmpcfg1'') else if (((b__0 = ( 0x3A2 :: 12 Word.word)))) then (''pmpcfg2'') else if (((b__0 = ( 0x3A3 :: 12 Word.word)))) then (''pmpcfg3'') else if (((b__0 = ( 0x3B0 :: 12 Word.word)))) then (''pmpaddr0'') else if (((b__0 = ( 0x3B1 :: 12 Word.word)))) then (''pmpaddr1'') else if (((b__0 = ( 0x3B2 :: 12 Word.word)))) then (''pmpaddr2'') else if (((b__0 = ( 0x3B3 :: 12 Word.word)))) then (''pmpaddr3'') else if (((b__0 = ( 0x3B4 :: 12 Word.word)))) then (''pmpaddr4'') else if (((b__0 = ( 0x3B5 :: 12 Word.word)))) then (''pmpaddr5'') else if (((b__0 = ( 0x3B6 :: 12 Word.word)))) then (''pmpaddr6'') else if (((b__0 = ( 0x3B7 :: 12 Word.word)))) then (''pmpaddr7'') else if (((b__0 = ( 0x3B8 :: 12 Word.word)))) then (''pmpaddr8'') else if (((b__0 = ( 0x3B9 :: 12 Word.word)))) then (''pmpaddr9'') else if (((b__0 = ( 0x3BA :: 12 Word.word)))) then (''pmpaddr10'') else if (((b__0 = ( 0x3BB :: 12 Word.word)))) then (''pmpaddr11'') else if (((b__0 = ( 0x3BC :: 12 Word.word)))) then (''pmpaddr12'') else if (((b__0 = ( 0x3BD :: 12 Word.word)))) then (''pmpaddr13'') else if (((b__0 = ( 0x3BE :: 12 Word.word)))) then (''pmpaddr14'') else if (((b__0 = ( 0x3BF :: 12 Word.word)))) then (''pmpaddr15'') else if (((b__0 = ( 0xB00 :: 12 Word.word)))) then (''mcycle'') else if (((b__0 = ( 0xB02 :: 12 Word.word)))) then (''minstret'') else if (((b__0 = ( 0xB80 :: 12 Word.word)))) then (''mcycleh'') else if (((b__0 = ( 0xB82 :: 12 Word.word)))) then (''minstreth'') else if (((b__0 = ( 0x7A0 :: 12 Word.word)))) then (''tselect'') else if (((b__0 = ( 0x7A1 :: 12 Word.word)))) then (''tdata1'') else if (((b__0 = ( 0x7A2 :: 12 Word.word)))) then (''tdata2'') else if (((b__0 = ( 0x7A3 :: 12 Word.word)))) then (''tdata3'') else decimal_string_of_bits b__0))\ for arg1 :: "(12)Word.word " definition csr_name_map_backwards :: \ string \((register_value),((12)Word.word),(exception))monad \ where \ csr_name_map_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''ustatus'')))) then return ( 0x000 :: 12 Word.word) else if (((p00 = (''uie'')))) then return ( 0x004 :: 12 Word.word) else if (((p00 = (''utvec'')))) then return ( 0x005 :: 12 Word.word) else if (((p00 = (''uscratch'')))) then return ( 0x040 :: 12 Word.word) else if (((p00 = (''uepc'')))) then return ( 0x041 :: 12 Word.word) else if (((p00 = (''ucause'')))) then return ( 0x042 :: 12 Word.word) else if (((p00 = (''utval'')))) then return ( 0x043 :: 12 Word.word) else if (((p00 = (''uip'')))) then return ( 0x044 :: 12 Word.word) else if (((p00 = (''fflags'')))) then return ( 0x001 :: 12 Word.word) else if (((p00 = (''frm'')))) then return ( 0x002 :: 12 Word.word) else if (((p00 = (''fcsr'')))) then return ( 0x003 :: 12 Word.word) else if (((p00 = (''cycle'')))) then return ( 0xC00 :: 12 Word.word) else if (((p00 = (''time'')))) then return ( 0xC01 :: 12 Word.word) else if (((p00 = (''instret'')))) then return ( 0xC02 :: 12 Word.word) else if (((p00 = (''cycleh'')))) then return ( 0xC80 :: 12 Word.word) else if (((p00 = (''timeh'')))) then return ( 0xC81 :: 12 Word.word) else if (((p00 = (''instreth'')))) then return ( 0xC82 :: 12 Word.word) else if (((p00 = (''sstatus'')))) then return ( 0x100 :: 12 Word.word) else if (((p00 = (''sedeleg'')))) then return ( 0x102 :: 12 Word.word) else if (((p00 = (''sideleg'')))) then return ( 0x103 :: 12 Word.word) else if (((p00 = (''sie'')))) then return ( 0x104 :: 12 Word.word) else if (((p00 = (''stvec'')))) then return ( 0x105 :: 12 Word.word) else if (((p00 = (''scounteren'')))) then return ( 0x106 :: 12 Word.word) else if (((p00 = (''sscratch'')))) then return ( 0x140 :: 12 Word.word) else if (((p00 = (''sepc'')))) then return ( 0x141 :: 12 Word.word) else if (((p00 = (''scause'')))) then return ( 0x142 :: 12 Word.word) else if (((p00 = (''stval'')))) then return ( 0x143 :: 12 Word.word) else if (((p00 = (''sip'')))) then return ( 0x144 :: 12 Word.word) else if (((p00 = (''satp'')))) then return ( 0x180 :: 12 Word.word) else if (((p00 = (''mvendorid'')))) then return ( 0xF11 :: 12 Word.word) else if (((p00 = (''marchid'')))) then return ( 0xF12 :: 12 Word.word) else if (((p00 = (''mimpid'')))) then return ( 0xF13 :: 12 Word.word) else if (((p00 = (''mhartid'')))) then return ( 0xF14 :: 12 Word.word) else if (((p00 = (''mstatus'')))) then return ( 0x300 :: 12 Word.word) else if (((p00 = (''misa'')))) then return ( 0x301 :: 12 Word.word) else if (((p00 = (''medeleg'')))) then return ( 0x302 :: 12 Word.word) else if (((p00 = (''mideleg'')))) then return ( 0x303 :: 12 Word.word) else if (((p00 = (''mie'')))) then return ( 0x304 :: 12 Word.word) else if (((p00 = (''mtvec'')))) then return ( 0x305 :: 12 Word.word) else if (((p00 = (''mcounteren'')))) then return ( 0x306 :: 12 Word.word) else if (((p00 = (''mcountinhibit'')))) then return ( 0x320 :: 12 Word.word) else if (((p00 = (''mscratch'')))) then return ( 0x340 :: 12 Word.word) else if (((p00 = (''mepc'')))) then return ( 0x341 :: 12 Word.word) else if (((p00 = (''mcause'')))) then return ( 0x342 :: 12 Word.word) else if (((p00 = (''mtval'')))) then return ( 0x343 :: 12 Word.word) else if (((p00 = (''mip'')))) then return ( 0x344 :: 12 Word.word) else if (((p00 = (''pmpcfg0'')))) then return ( 0x3A0 :: 12 Word.word) else if (((p00 = (''pmpcfg1'')))) then return ( 0x3A1 :: 12 Word.word) else if (((p00 = (''pmpcfg2'')))) then return ( 0x3A2 :: 12 Word.word) else if (((p00 = (''pmpcfg3'')))) then return ( 0x3A3 :: 12 Word.word) else if (((p00 = (''pmpaddr0'')))) then return ( 0x3B0 :: 12 Word.word) else if (((p00 = (''pmpaddr1'')))) then return ( 0x3B1 :: 12 Word.word) else if (((p00 = (''pmpaddr2'')))) then return ( 0x3B2 :: 12 Word.word) else if (((p00 = (''pmpaddr3'')))) then return ( 0x3B3 :: 12 Word.word) else if (((p00 = (''pmpaddr4'')))) then return ( 0x3B4 :: 12 Word.word) else if (((p00 = (''pmpaddr5'')))) then return ( 0x3B5 :: 12 Word.word) else if (((p00 = (''pmpaddr6'')))) then return ( 0x3B6 :: 12 Word.word) else if (((p00 = (''pmpaddr7'')))) then return ( 0x3B7 :: 12 Word.word) else if (((p00 = (''pmpaddr8'')))) then return ( 0x3B8 :: 12 Word.word) else if (((p00 = (''pmpaddr9'')))) then return ( 0x3B9 :: 12 Word.word) else if (((p00 = (''pmpaddr10'')))) then return ( 0x3BA :: 12 Word.word) else if (((p00 = (''pmpaddr11'')))) then return ( 0x3BB :: 12 Word.word) else if (((p00 = (''pmpaddr12'')))) then return ( 0x3BC :: 12 Word.word) else if (((p00 = (''pmpaddr13'')))) then return ( 0x3BD :: 12 Word.word) else if (((p00 = (''pmpaddr14'')))) then return ( 0x3BE :: 12 Word.word) else if (((p00 = (''pmpaddr15'')))) then return ( 0x3BF :: 12 Word.word) else if (((p00 = (''mcycle'')))) then return ( 0xB00 :: 12 Word.word) else if (((p00 = (''minstret'')))) then return ( 0xB02 :: 12 Word.word) else if (((p00 = (''mcycleh'')))) then return ( 0xB80 :: 12 Word.word) else if (((p00 = (''minstreth'')))) then return ( 0xB82 :: 12 Word.word) else if (((p00 = (''tselect'')))) then return ( 0x7A0 :: 12 Word.word) else if (((p00 = (''tdata1'')))) then return ( 0x7A1 :: 12 Word.word) else if (((p00 = (''tdata2'')))) then return ( 0x7A2 :: 12 Word.word) else if (((p00 = (''tdata3'')))) then return ( 0x7A3 :: 12 Word.word) else if ((hex_bits_12_backwards_matches p00)) then (hex_bits_12_backwards p00 :: ( 12 Word.word) M) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition csr_name_map_forwards_matches :: \(12)Word.word \ bool \ where \ csr_name_map_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0x000 :: 12 Word.word)))) then True else if (((b__0 = ( 0x004 :: 12 Word.word)))) then True else if (((b__0 = ( 0x005 :: 12 Word.word)))) then True else if (((b__0 = ( 0x040 :: 12 Word.word)))) then True else if (((b__0 = ( 0x041 :: 12 Word.word)))) then True else if (((b__0 = ( 0x042 :: 12 Word.word)))) then True else if (((b__0 = ( 0x043 :: 12 Word.word)))) then True else if (((b__0 = ( 0x044 :: 12 Word.word)))) then True else if (((b__0 = ( 0x001 :: 12 Word.word)))) then True else if (((b__0 = ( 0x002 :: 12 Word.word)))) then True else if (((b__0 = ( 0x003 :: 12 Word.word)))) then True else if (((b__0 = ( 0xC00 :: 12 Word.word)))) then True else if (((b__0 = ( 0xC01 :: 12 Word.word)))) then True else if (((b__0 = ( 0xC02 :: 12 Word.word)))) then True else if (((b__0 = ( 0xC80 :: 12 Word.word)))) then True else if (((b__0 = ( 0xC81 :: 12 Word.word)))) then True else if (((b__0 = ( 0xC82 :: 12 Word.word)))) then True else if (((b__0 = ( 0x100 :: 12 Word.word)))) then True else if (((b__0 = ( 0x102 :: 12 Word.word)))) then True else if (((b__0 = ( 0x103 :: 12 Word.word)))) then True else if (((b__0 = ( 0x104 :: 12 Word.word)))) then True else if (((b__0 = ( 0x105 :: 12 Word.word)))) then True else if (((b__0 = ( 0x106 :: 12 Word.word)))) then True else if (((b__0 = ( 0x140 :: 12 Word.word)))) then True else if (((b__0 = ( 0x141 :: 12 Word.word)))) then True else if (((b__0 = ( 0x142 :: 12 Word.word)))) then True else if (((b__0 = ( 0x143 :: 12 Word.word)))) then True else if (((b__0 = ( 0x144 :: 12 Word.word)))) then True else if (((b__0 = ( 0x180 :: 12 Word.word)))) then True else if (((b__0 = ( 0xF11 :: 12 Word.word)))) then True else if (((b__0 = ( 0xF12 :: 12 Word.word)))) then True else if (((b__0 = ( 0xF13 :: 12 Word.word)))) then True else if (((b__0 = ( 0xF14 :: 12 Word.word)))) then True else if (((b__0 = ( 0x300 :: 12 Word.word)))) then True else if (((b__0 = ( 0x301 :: 12 Word.word)))) then True else if (((b__0 = ( 0x302 :: 12 Word.word)))) then True else if (((b__0 = ( 0x303 :: 12 Word.word)))) then True else if (((b__0 = ( 0x304 :: 12 Word.word)))) then True else if (((b__0 = ( 0x305 :: 12 Word.word)))) then True else if (((b__0 = ( 0x306 :: 12 Word.word)))) then True else if (((b__0 = ( 0x320 :: 12 Word.word)))) then True else if (((b__0 = ( 0x340 :: 12 Word.word)))) then True else if (((b__0 = ( 0x341 :: 12 Word.word)))) then True else if (((b__0 = ( 0x342 :: 12 Word.word)))) then True else if (((b__0 = ( 0x343 :: 12 Word.word)))) then True else if (((b__0 = ( 0x344 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3A0 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3A1 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3A2 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3A3 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B0 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B1 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B2 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B3 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B4 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B5 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B6 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B7 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B8 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3B9 :: 12 Word.word)))) then True else if (((b__0 = ( 0x3BA :: 12 Word.word)))) then True else if (((b__0 = ( 0x3BB :: 12 Word.word)))) then True else if (((b__0 = ( 0x3BC :: 12 Word.word)))) then True else if (((b__0 = ( 0x3BD :: 12 Word.word)))) then True else if (((b__0 = ( 0x3BE :: 12 Word.word)))) then True else if (((b__0 = ( 0x3BF :: 12 Word.word)))) then True else if (((b__0 = ( 0xB00 :: 12 Word.word)))) then True else if (((b__0 = ( 0xB02 :: 12 Word.word)))) then True else if (((b__0 = ( 0xB80 :: 12 Word.word)))) then True else if (((b__0 = ( 0xB82 :: 12 Word.word)))) then True else if (((b__0 = ( 0x7A0 :: 12 Word.word)))) then True else if (((b__0 = ( 0x7A1 :: 12 Word.word)))) then True else if (((b__0 = ( 0x7A2 :: 12 Word.word)))) then True else if (((b__0 = ( 0x7A3 :: 12 Word.word)))) then True else True))\ for arg1 :: "(12)Word.word " definition csr_name_map_backwards_matches :: \ string \((register_value),(bool),(exception))monad \ where \ csr_name_map_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''ustatus'')))) then return True else if (((p00 = (''uie'')))) then return True else if (((p00 = (''utvec'')))) then return True else if (((p00 = (''uscratch'')))) then return True else if (((p00 = (''uepc'')))) then return True else if (((p00 = (''ucause'')))) then return True else if (((p00 = (''utval'')))) then return True else if (((p00 = (''uip'')))) then return True else if (((p00 = (''fflags'')))) then return True else if (((p00 = (''frm'')))) then return True else if (((p00 = (''fcsr'')))) then return True else if (((p00 = (''cycle'')))) then return True else if (((p00 = (''time'')))) then return True else if (((p00 = (''instret'')))) then return True else if (((p00 = (''cycleh'')))) then return True else if (((p00 = (''timeh'')))) then return True else if (((p00 = (''instreth'')))) then return True else if (((p00 = (''sstatus'')))) then return True else if (((p00 = (''sedeleg'')))) then return True else if (((p00 = (''sideleg'')))) then return True else if (((p00 = (''sie'')))) then return True else if (((p00 = (''stvec'')))) then return True else if (((p00 = (''scounteren'')))) then return True else if (((p00 = (''sscratch'')))) then return True else if (((p00 = (''sepc'')))) then return True else if (((p00 = (''scause'')))) then return True else if (((p00 = (''stval'')))) then return True else if (((p00 = (''sip'')))) then return True else if (((p00 = (''satp'')))) then return True else if (((p00 = (''mvendorid'')))) then return True else if (((p00 = (''marchid'')))) then return True else if (((p00 = (''mimpid'')))) then return True else if (((p00 = (''mhartid'')))) then return True else if (((p00 = (''mstatus'')))) then return True else if (((p00 = (''misa'')))) then return True else if (((p00 = (''medeleg'')))) then return True else if (((p00 = (''mideleg'')))) then return True else if (((p00 = (''mie'')))) then return True else if (((p00 = (''mtvec'')))) then return True else if (((p00 = (''mcounteren'')))) then return True else if (((p00 = (''mcountinhibit'')))) then return True else if (((p00 = (''mscratch'')))) then return True else if (((p00 = (''mepc'')))) then return True else if (((p00 = (''mcause'')))) then return True else if (((p00 = (''mtval'')))) then return True else if (((p00 = (''mip'')))) then return True else if (((p00 = (''pmpcfg0'')))) then return True else if (((p00 = (''pmpcfg1'')))) then return True else if (((p00 = (''pmpcfg2'')))) then return True else if (((p00 = (''pmpcfg3'')))) then return True else if (((p00 = (''pmpaddr0'')))) then return True else if (((p00 = (''pmpaddr1'')))) then return True else if (((p00 = (''pmpaddr2'')))) then return True else if (((p00 = (''pmpaddr3'')))) then return True else if (((p00 = (''pmpaddr4'')))) then return True else if (((p00 = (''pmpaddr5'')))) then return True else if (((p00 = (''pmpaddr6'')))) then return True else if (((p00 = (''pmpaddr7'')))) then return True else if (((p00 = (''pmpaddr8'')))) then return True else if (((p00 = (''pmpaddr9'')))) then return True else if (((p00 = (''pmpaddr10'')))) then return True else if (((p00 = (''pmpaddr11'')))) then return True else if (((p00 = (''pmpaddr12'')))) then return True else if (((p00 = (''pmpaddr13'')))) then return True else if (((p00 = (''pmpaddr14'')))) then return True else if (((p00 = (''pmpaddr15'')))) then return True else if (((p00 = (''mcycle'')))) then return True else if (((p00 = (''minstret'')))) then return True else if (((p00 = (''mcycleh'')))) then return True else if (((p00 = (''minstreth'')))) then return True else if (((p00 = (''tselect'')))) then return True else if (((p00 = (''tdata1'')))) then return True else if (((p00 = (''tdata2'')))) then return True else if (((p00 = (''tdata3'')))) then return True else if ((hex_bits_12_backwards_matches p00)) then (hex_bits_12_backwards p00 :: ( 12 Word.word) M) \ ((\ reg . return True)) else return False))\ for arg1 :: " string " \ \\val _s752_ : string -> maybe ((mword ty12 * string))\\ definition s752 :: \ string \((12)Word.word*string)option \ where \ s752 s7540 = ( (case ((hex_bits_12_matches_prefix0 s7540 :: (( 12 Word.word * ii))option)) of Some ((reg, s7550)) => (case ((string_drop s7540 s7550)) of s1 => Some (reg, s1) ) | _ => None ))\ for s7540 :: " string " \ \\val _s748_ : string -> maybe string\\ definition s748 :: \ string \(string)option \ where \ s748 s7490 = ( (let s7500 = s7490 in if ((string_startswith s7500 (''tdata3''))) then (case ((string_drop s7500 ((string_length (''tdata3''))))) of s1 => Some s1 ) else None))\ for s7490 :: " string " \ \\val _s744_ : string -> maybe string\\ definition s744 :: \ string \(string)option \ where \ s744 s7450 = ( (let s7460 = s7450 in if ((string_startswith s7460 (''tdata2''))) then (case ((string_drop s7460 ((string_length (''tdata2''))))) of s1 => Some s1 ) else None))\ for s7450 :: " string " \ \\val _s740_ : string -> maybe string\\ definition s740 :: \ string \(string)option \ where \ s740 s7410 = ( (let s7420 = s7410 in if ((string_startswith s7420 (''tdata1''))) then (case ((string_drop s7420 ((string_length (''tdata1''))))) of s1 => Some s1 ) else None))\ for s7410 :: " string " \ \\val _s736_ : string -> maybe string\\ definition s736 :: \ string \(string)option \ where \ s736 s7370 = ( (let s7380 = s7370 in if ((string_startswith s7380 (''tselect''))) then (case ((string_drop s7380 ((string_length (''tselect''))))) of s1 => Some s1 ) else None))\ for s7370 :: " string " \ \\val _s732_ : string -> maybe string\\ definition s732 :: \ string \(string)option \ where \ s732 s7330 = ( (let s7340 = s7330 in if ((string_startswith s7340 (''minstreth''))) then (case ((string_drop s7340 ((string_length (''minstreth''))))) of s1 => Some s1 ) else None))\ for s7330 :: " string " \ \\val _s728_ : string -> maybe string\\ definition s728 :: \ string \(string)option \ where \ s728 s7290 = ( (let s7300 = s7290 in if ((string_startswith s7300 (''mcycleh''))) then (case ((string_drop s7300 ((string_length (''mcycleh''))))) of s1 => Some s1 ) else None))\ for s7290 :: " string " \ \\val _s724_ : string -> maybe string\\ definition s724 :: \ string \(string)option \ where \ s724 s7250 = ( (let s7260 = s7250 in if ((string_startswith s7260 (''minstret''))) then (case ((string_drop s7260 ((string_length (''minstret''))))) of s1 => Some s1 ) else None))\ for s7250 :: " string " \ \\val _s720_ : string -> maybe string\\ definition s720 :: \ string \(string)option \ where \ s720 s7210 = ( (let s7220 = s7210 in if ((string_startswith s7220 (''mcycle''))) then (case ((string_drop s7220 ((string_length (''mcycle''))))) of s1 => Some s1 ) else None))\ for s7210 :: " string " \ \\val _s716_ : string -> maybe string\\ definition s716 :: \ string \(string)option \ where \ s716 s7170 = ( (let s7180 = s7170 in if ((string_startswith s7180 (''pmpaddr15''))) then (case ((string_drop s7180 ((string_length (''pmpaddr15''))))) of s1 => Some s1 ) else None))\ for s7170 :: " string " \ \\val _s712_ : string -> maybe string\\ definition s712 :: \ string \(string)option \ where \ s712 s7130 = ( (let s7140 = s7130 in if ((string_startswith s7140 (''pmpaddr14''))) then (case ((string_drop s7140 ((string_length (''pmpaddr14''))))) of s1 => Some s1 ) else None))\ for s7130 :: " string " \ \\val _s708_ : string -> maybe string\\ definition s708 :: \ string \(string)option \ where \ s708 s7090 = ( (let s7100 = s7090 in if ((string_startswith s7100 (''pmpaddr13''))) then (case ((string_drop s7100 ((string_length (''pmpaddr13''))))) of s1 => Some s1 ) else None))\ for s7090 :: " string " \ \\val _s704_ : string -> maybe string\\ definition s704 :: \ string \(string)option \ where \ s704 s7050 = ( (let s7060 = s7050 in if ((string_startswith s7060 (''pmpaddr12''))) then (case ((string_drop s7060 ((string_length (''pmpaddr12''))))) of s1 => Some s1 ) else None))\ for s7050 :: " string " \ \\val _s700_ : string -> maybe string\\ definition s700 :: \ string \(string)option \ where \ s700 s7010 = ( (let s7020 = s7010 in if ((string_startswith s7020 (''pmpaddr11''))) then (case ((string_drop s7020 ((string_length (''pmpaddr11''))))) of s1 => Some s1 ) else None))\ for s7010 :: " string " \ \\val _s696_ : string -> maybe string\\ definition s696 :: \ string \(string)option \ where \ s696 s6970 = ( (let s6980 = s6970 in if ((string_startswith s6980 (''pmpaddr10''))) then (case ((string_drop s6980 ((string_length (''pmpaddr10''))))) of s1 => Some s1 ) else None))\ for s6970 :: " string " \ \\val _s692_ : string -> maybe string\\ definition s692 :: \ string \(string)option \ where \ s692 s6930 = ( (let s6940 = s6930 in if ((string_startswith s6940 (''pmpaddr9''))) then (case ((string_drop s6940 ((string_length (''pmpaddr9''))))) of s1 => Some s1 ) else None))\ for s6930 :: " string " \ \\val _s688_ : string -> maybe string\\ definition s688 :: \ string \(string)option \ where \ s688 s6890 = ( (let s6900 = s6890 in if ((string_startswith s6900 (''pmpaddr8''))) then (case ((string_drop s6900 ((string_length (''pmpaddr8''))))) of s1 => Some s1 ) else None))\ for s6890 :: " string " \ \\val _s684_ : string -> maybe string\\ definition s684 :: \ string \(string)option \ where \ s684 s6850 = ( (let s6860 = s6850 in if ((string_startswith s6860 (''pmpaddr7''))) then (case ((string_drop s6860 ((string_length (''pmpaddr7''))))) of s1 => Some s1 ) else None))\ for s6850 :: " string " \ \\val _s680_ : string -> maybe string\\ definition s680 :: \ string \(string)option \ where \ s680 s6810 = ( (let s6820 = s6810 in if ((string_startswith s6820 (''pmpaddr6''))) then (case ((string_drop s6820 ((string_length (''pmpaddr6''))))) of s1 => Some s1 ) else None))\ for s6810 :: " string " \ \\val _s676_ : string -> maybe string\\ definition s676 :: \ string \(string)option \ where \ s676 s6770 = ( (let s6780 = s6770 in if ((string_startswith s6780 (''pmpaddr5''))) then (case ((string_drop s6780 ((string_length (''pmpaddr5''))))) of s1 => Some s1 ) else None))\ for s6770 :: " string " \ \\val _s672_ : string -> maybe string\\ definition s672 :: \ string \(string)option \ where \ s672 s6730 = ( (let s6740 = s6730 in if ((string_startswith s6740 (''pmpaddr4''))) then (case ((string_drop s6740 ((string_length (''pmpaddr4''))))) of s1 => Some s1 ) else None))\ for s6730 :: " string " \ \\val _s668_ : string -> maybe string\\ definition s668 :: \ string \(string)option \ where \ s668 s6690 = ( (let s6700 = s6690 in if ((string_startswith s6700 (''pmpaddr3''))) then (case ((string_drop s6700 ((string_length (''pmpaddr3''))))) of s1 => Some s1 ) else None))\ for s6690 :: " string " \ \\val _s664_ : string -> maybe string\\ definition s664 :: \ string \(string)option \ where \ s664 s6650 = ( (let s6660 = s6650 in if ((string_startswith s6660 (''pmpaddr2''))) then (case ((string_drop s6660 ((string_length (''pmpaddr2''))))) of s1 => Some s1 ) else None))\ for s6650 :: " string " \ \\val _s660_ : string -> maybe string\\ definition s660 :: \ string \(string)option \ where \ s660 s6610 = ( (let s6620 = s6610 in if ((string_startswith s6620 (''pmpaddr1''))) then (case ((string_drop s6620 ((string_length (''pmpaddr1''))))) of s1 => Some s1 ) else None))\ for s6610 :: " string " \ \\val _s656_ : string -> maybe string\\ definition s656 :: \ string \(string)option \ where \ s656 s6570 = ( (let s6580 = s6570 in if ((string_startswith s6580 (''pmpaddr0''))) then (case ((string_drop s6580 ((string_length (''pmpaddr0''))))) of s1 => Some s1 ) else None))\ for s6570 :: " string " \ \\val _s652_ : string -> maybe string\\ definition s652 :: \ string \(string)option \ where \ s652 s6530 = ( (let s6540 = s6530 in if ((string_startswith s6540 (''pmpcfg3''))) then (case ((string_drop s6540 ((string_length (''pmpcfg3''))))) of s1 => Some s1 ) else None))\ for s6530 :: " string " \ \\val _s648_ : string -> maybe string\\ definition s648 :: \ string \(string)option \ where \ s648 s6490 = ( (let s6500 = s6490 in if ((string_startswith s6500 (''pmpcfg2''))) then (case ((string_drop s6500 ((string_length (''pmpcfg2''))))) of s1 => Some s1 ) else None))\ for s6490 :: " string " \ \\val _s644_ : string -> maybe string\\ definition s644 :: \ string \(string)option \ where \ s644 s6450 = ( (let s6460 = s6450 in if ((string_startswith s6460 (''pmpcfg1''))) then (case ((string_drop s6460 ((string_length (''pmpcfg1''))))) of s1 => Some s1 ) else None))\ for s6450 :: " string " \ \\val _s640_ : string -> maybe string\\ definition s640 :: \ string \(string)option \ where \ s640 s6410 = ( (let s6420 = s6410 in if ((string_startswith s6420 (''pmpcfg0''))) then (case ((string_drop s6420 ((string_length (''pmpcfg0''))))) of s1 => Some s1 ) else None))\ for s6410 :: " string " \ \\val _s636_ : string -> maybe string\\ definition s636 :: \ string \(string)option \ where \ s636 s6370 = ( (let s6380 = s6370 in if ((string_startswith s6380 (''mip''))) then (case ((string_drop s6380 ((string_length (''mip''))))) of s1 => Some s1 ) else None))\ for s6370 :: " string " \ \\val _s632_ : string -> maybe string\\ definition s632 :: \ string \(string)option \ where \ s632 s6330 = ( (let s6340 = s6330 in if ((string_startswith s6340 (''mtval''))) then (case ((string_drop s6340 ((string_length (''mtval''))))) of s1 => Some s1 ) else None))\ for s6330 :: " string " \ \\val _s628_ : string -> maybe string\\ definition s628 :: \ string \(string)option \ where \ s628 s6290 = ( (let s6300 = s6290 in if ((string_startswith s6300 (''mcause''))) then (case ((string_drop s6300 ((string_length (''mcause''))))) of s1 => Some s1 ) else None))\ for s6290 :: " string " \ \\val _s624_ : string -> maybe string\\ definition s624 :: \ string \(string)option \ where \ s624 s6250 = ( (let s6260 = s6250 in if ((string_startswith s6260 (''mepc''))) then (case ((string_drop s6260 ((string_length (''mepc''))))) of s1 => Some s1 ) else None))\ for s6250 :: " string " \ \\val _s620_ : string -> maybe string\\ definition s620 :: \ string \(string)option \ where \ s620 s6210 = ( (let s6220 = s6210 in if ((string_startswith s6220 (''mscratch''))) then (case ((string_drop s6220 ((string_length (''mscratch''))))) of s1 => Some s1 ) else None))\ for s6210 :: " string " \ \\val _s616_ : string -> maybe string\\ definition s616 :: \ string \(string)option \ where \ s616 s6170 = ( (let s6180 = s6170 in if ((string_startswith s6180 (''mcountinhibit''))) then (case ((string_drop s6180 ((string_length (''mcountinhibit''))))) of s1 => Some s1 ) else None))\ for s6170 :: " string " \ \\val _s612_ : string -> maybe string\\ definition s612 :: \ string \(string)option \ where \ s612 s6130 = ( (let s6140 = s6130 in if ((string_startswith s6140 (''mcounteren''))) then (case ((string_drop s6140 ((string_length (''mcounteren''))))) of s1 => Some s1 ) else None))\ for s6130 :: " string " \ \\val _s608_ : string -> maybe string\\ definition s608 :: \ string \(string)option \ where \ s608 s6090 = ( (let s6100 = s6090 in if ((string_startswith s6100 (''mtvec''))) then (case ((string_drop s6100 ((string_length (''mtvec''))))) of s1 => Some s1 ) else None))\ for s6090 :: " string " \ \\val _s604_ : string -> maybe string\\ definition s604 :: \ string \(string)option \ where \ s604 s6050 = ( (let s6060 = s6050 in if ((string_startswith s6060 (''mie''))) then (case ((string_drop s6060 ((string_length (''mie''))))) of s1 => Some s1 ) else None))\ for s6050 :: " string " \ \\val _s600_ : string -> maybe string\\ definition s600 :: \ string \(string)option \ where \ s600 s6010 = ( (let s6020 = s6010 in if ((string_startswith s6020 (''mideleg''))) then (case ((string_drop s6020 ((string_length (''mideleg''))))) of s1 => Some s1 ) else None))\ for s6010 :: " string " \ \\val _s596_ : string -> maybe string\\ definition s596 :: \ string \(string)option \ where \ s596 s5970 = ( (let s5980 = s5970 in if ((string_startswith s5980 (''medeleg''))) then (case ((string_drop s5980 ((string_length (''medeleg''))))) of s1 => Some s1 ) else None))\ for s5970 :: " string " \ \\val _s592_ : string -> maybe string\\ definition s592 :: \ string \(string)option \ where \ s592 s5930 = ( (let s5940 = s5930 in if ((string_startswith s5940 (''misa''))) then (case ((string_drop s5940 ((string_length (''misa''))))) of s1 => Some s1 ) else None))\ for s5930 :: " string " \ \\val _s588_ : string -> maybe string\\ definition s588 :: \ string \(string)option \ where \ s588 s5890 = ( (let s5900 = s5890 in if ((string_startswith s5900 (''mstatus''))) then (case ((string_drop s5900 ((string_length (''mstatus''))))) of s1 => Some s1 ) else None))\ for s5890 :: " string " \ \\val _s584_ : string -> maybe string\\ definition s584 :: \ string \(string)option \ where \ s584 s5850 = ( (let s5860 = s5850 in if ((string_startswith s5860 (''mhartid''))) then (case ((string_drop s5860 ((string_length (''mhartid''))))) of s1 => Some s1 ) else None))\ for s5850 :: " string " \ \\val _s580_ : string -> maybe string\\ definition s580 :: \ string \(string)option \ where \ s580 s5810 = ( (let s5820 = s5810 in if ((string_startswith s5820 (''mimpid''))) then (case ((string_drop s5820 ((string_length (''mimpid''))))) of s1 => Some s1 ) else None))\ for s5810 :: " string " \ \\val _s576_ : string -> maybe string\\ definition s576 :: \ string \(string)option \ where \ s576 s5770 = ( (let s5780 = s5770 in if ((string_startswith s5780 (''marchid''))) then (case ((string_drop s5780 ((string_length (''marchid''))))) of s1 => Some s1 ) else None))\ for s5770 :: " string " \ \\val _s572_ : string -> maybe string\\ definition s572 :: \ string \(string)option \ where \ s572 s5730 = ( (let s5740 = s5730 in if ((string_startswith s5740 (''mvendorid''))) then (case ((string_drop s5740 ((string_length (''mvendorid''))))) of s1 => Some s1 ) else None))\ for s5730 :: " string " \ \\val _s568_ : string -> maybe string\\ definition s568 :: \ string \(string)option \ where \ s568 s5690 = ( (let s5700 = s5690 in if ((string_startswith s5700 (''satp''))) then (case ((string_drop s5700 ((string_length (''satp''))))) of s1 => Some s1 ) else None))\ for s5690 :: " string " \ \\val _s564_ : string -> maybe string\\ definition s564 :: \ string \(string)option \ where \ s564 s5650 = ( (let s5660 = s5650 in if ((string_startswith s5660 (''sip''))) then (case ((string_drop s5660 ((string_length (''sip''))))) of s1 => Some s1 ) else None))\ for s5650 :: " string " \ \\val _s560_ : string -> maybe string\\ definition s560 :: \ string \(string)option \ where \ s560 s5610 = ( (let s5620 = s5610 in if ((string_startswith s5620 (''stval''))) then (case ((string_drop s5620 ((string_length (''stval''))))) of s1 => Some s1 ) else None))\ for s5610 :: " string " \ \\val _s556_ : string -> maybe string\\ definition s556 :: \ string \(string)option \ where \ s556 s5570 = ( (let s5580 = s5570 in if ((string_startswith s5580 (''scause''))) then (case ((string_drop s5580 ((string_length (''scause''))))) of s1 => Some s1 ) else None))\ for s5570 :: " string " \ \\val _s552_ : string -> maybe string\\ definition s552 :: \ string \(string)option \ where \ s552 s5530 = ( (let s5540 = s5530 in if ((string_startswith s5540 (''sepc''))) then (case ((string_drop s5540 ((string_length (''sepc''))))) of s1 => Some s1 ) else None))\ for s5530 :: " string " \ \\val _s548_ : string -> maybe string\\ definition s548 :: \ string \(string)option \ where \ s548 s5490 = ( (let s5500 = s5490 in if ((string_startswith s5500 (''sscratch''))) then (case ((string_drop s5500 ((string_length (''sscratch''))))) of s1 => Some s1 ) else None))\ for s5490 :: " string " \ \\val _s544_ : string -> maybe string\\ definition s544 :: \ string \(string)option \ where \ s544 s5450 = ( (let s5460 = s5450 in if ((string_startswith s5460 (''scounteren''))) then (case ((string_drop s5460 ((string_length (''scounteren''))))) of s1 => Some s1 ) else None))\ for s5450 :: " string " \ \\val _s540_ : string -> maybe string\\ definition s540 :: \ string \(string)option \ where \ s540 s5410 = ( (let s5420 = s5410 in if ((string_startswith s5420 (''stvec''))) then (case ((string_drop s5420 ((string_length (''stvec''))))) of s1 => Some s1 ) else None))\ for s5410 :: " string " \ \\val _s536_ : string -> maybe string\\ definition s536 :: \ string \(string)option \ where \ s536 s5370 = ( (let s5380 = s5370 in if ((string_startswith s5380 (''sie''))) then (case ((string_drop s5380 ((string_length (''sie''))))) of s1 => Some s1 ) else None))\ for s5370 :: " string " \ \\val _s532_ : string -> maybe string\\ definition s532 :: \ string \(string)option \ where \ s532 s5330 = ( (let s5340 = s5330 in if ((string_startswith s5340 (''sideleg''))) then (case ((string_drop s5340 ((string_length (''sideleg''))))) of s1 => Some s1 ) else None))\ for s5330 :: " string " \ \\val _s528_ : string -> maybe string\\ definition s528 :: \ string \(string)option \ where \ s528 s5290 = ( (let s5300 = s5290 in if ((string_startswith s5300 (''sedeleg''))) then (case ((string_drop s5300 ((string_length (''sedeleg''))))) of s1 => Some s1 ) else None))\ for s5290 :: " string " \ \\val _s524_ : string -> maybe string\\ definition s524 :: \ string \(string)option \ where \ s524 s5250 = ( (let s5260 = s5250 in if ((string_startswith s5260 (''sstatus''))) then (case ((string_drop s5260 ((string_length (''sstatus''))))) of s1 => Some s1 ) else None))\ for s5250 :: " string " \ \\val _s520_ : string -> maybe string\\ definition s520 :: \ string \(string)option \ where \ s520 s5210 = ( (let s5220 = s5210 in if ((string_startswith s5220 (''instreth''))) then (case ((string_drop s5220 ((string_length (''instreth''))))) of s1 => Some s1 ) else None))\ for s5210 :: " string " \ \\val _s516_ : string -> maybe string\\ definition s516 :: \ string \(string)option \ where \ s516 s5170 = ( (let s5180 = s5170 in if ((string_startswith s5180 (''timeh''))) then (case ((string_drop s5180 ((string_length (''timeh''))))) of s1 => Some s1 ) else None))\ for s5170 :: " string " \ \\val _s512_ : string -> maybe string\\ definition s512 :: \ string \(string)option \ where \ s512 s5130 = ( (let s5140 = s5130 in if ((string_startswith s5140 (''cycleh''))) then (case ((string_drop s5140 ((string_length (''cycleh''))))) of s1 => Some s1 ) else None))\ for s5130 :: " string " \ \\val _s508_ : string -> maybe string\\ definition s508 :: \ string \(string)option \ where \ s508 s5090 = ( (let s5100 = s5090 in if ((string_startswith s5100 (''instret''))) then (case ((string_drop s5100 ((string_length (''instret''))))) of s1 => Some s1 ) else None))\ for s5090 :: " string " \ \\val _s504_ : string -> maybe string\\ definition s504 :: \ string \(string)option \ where \ s504 s5050 = ( (let s5060 = s5050 in if ((string_startswith s5060 (''time''))) then (case ((string_drop s5060 ((string_length (''time''))))) of s1 => Some s1 ) else None))\ for s5050 :: " string " \ \\val _s500_ : string -> maybe string\\ definition s500 :: \ string \(string)option \ where \ s500 s5011 = ( (let s5020 = s5011 in if ((string_startswith s5020 (''cycle''))) then (case ((string_drop s5020 ((string_length (''cycle''))))) of s1 => Some s1 ) else None))\ for s5011 :: " string " \ \\val _s496_ : string -> maybe string\\ definition s496 :: \ string \(string)option \ where \ s496 s4970 = ( (let s4980 = s4970 in if ((string_startswith s4980 (''fcsr''))) then (case ((string_drop s4980 ((string_length (''fcsr''))))) of s1 => Some s1 ) else None))\ for s4970 :: " string " \ \\val _s492_ : string -> maybe string\\ definition s492 :: \ string \(string)option \ where \ s492 s4930 = ( (let s4940 = s4930 in if ((string_startswith s4940 (''frm''))) then (case ((string_drop s4940 ((string_length (''frm''))))) of s1 => Some s1 ) else None))\ for s4930 :: " string " \ \\val _s488_ : string -> maybe string\\ definition s488 :: \ string \(string)option \ where \ s488 s4890 = ( (let s4900 = s4890 in if ((string_startswith s4900 (''fflags''))) then (case ((string_drop s4900 ((string_length (''fflags''))))) of s1 => Some s1 ) else None))\ for s4890 :: " string " \ \\val _s484_ : string -> maybe string\\ definition s484 :: \ string \(string)option \ where \ s484 s4850 = ( (let s4860 = s4850 in if ((string_startswith s4860 (''uip''))) then (case ((string_drop s4860 ((string_length (''uip''))))) of s1 => Some s1 ) else None))\ for s4850 :: " string " \ \\val _s480_ : string -> maybe string\\ definition s480 :: \ string \(string)option \ where \ s480 s4810 = ( (let s4820 = s4810 in if ((string_startswith s4820 (''utval''))) then (case ((string_drop s4820 ((string_length (''utval''))))) of s1 => Some s1 ) else None))\ for s4810 :: " string " \ \\val _s476_ : string -> maybe string\\ definition s476 :: \ string \(string)option \ where \ s476 s4770 = ( (let s4780 = s4770 in if ((string_startswith s4780 (''ucause''))) then (case ((string_drop s4780 ((string_length (''ucause''))))) of s1 => Some s1 ) else None))\ for s4770 :: " string " \ \\val _s472_ : string -> maybe string\\ definition s472 :: \ string \(string)option \ where \ s472 s4730 = ( (let s4740 = s4730 in if ((string_startswith s4740 (''uepc''))) then (case ((string_drop s4740 ((string_length (''uepc''))))) of s1 => Some s1 ) else None))\ for s4730 :: " string " \ \\val _s468_ : string -> maybe string\\ definition s468 :: \ string \(string)option \ where \ s468 s4690 = ( (let s4700 = s4690 in if ((string_startswith s4700 (''uscratch''))) then (case ((string_drop s4700 ((string_length (''uscratch''))))) of s1 => Some s1 ) else None))\ for s4690 :: " string " \ \\val _s464_ : string -> maybe string\\ definition s464 :: \ string \(string)option \ where \ s464 s4650 = ( (let s4660 = s4650 in if ((string_startswith s4660 (''utvec''))) then (case ((string_drop s4660 ((string_length (''utvec''))))) of s1 => Some s1 ) else None))\ for s4650 :: " string " \ \\val _s460_ : string -> maybe string\\ definition s460 :: \ string \(string)option \ where \ s460 s4610 = ( (let s4621 = s4610 in if ((string_startswith s4621 (''uie''))) then (case ((string_drop s4621 ((string_length (''uie''))))) of s1 => Some s1 ) else None))\ for s4610 :: " string " \ \\val _s456_ : string -> maybe string\\ definition s456 :: \ string \(string)option \ where \ s456 s4570 = ( (let s4580 = s4570 in if ((string_startswith s4580 (''ustatus''))) then (case ((string_drop s4580 ((string_length (''ustatus''))))) of s1 => Some s1 ) else None))\ for s4570 :: " string " definition csr_name_map_matches_prefix :: \ string \((12)Word.word*int)option \ where \ csr_name_map_matches_prefix arg1 = ( (let s4590 = arg1 in if ((case ((s456 s4590)) of Some (s1) => True | _ => False )) then (case s456 s4590 of (Some (s1)) => Some (( 0x000 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s460 s4590)) of Some (s1) => True | _ => False )) then (case s460 s4590 of (Some (s1)) => Some (( 0x004 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s464 s4590)) of Some (s1) => True | _ => False )) then (case s464 s4590 of (Some (s1)) => Some (( 0x005 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s468 s4590)) of Some (s1) => True | _ => False )) then (case s468 s4590 of (Some (s1)) => Some (( 0x040 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s472 s4590)) of Some (s1) => True | _ => False )) then (case s472 s4590 of (Some (s1)) => Some (( 0x041 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s476 s4590)) of Some (s1) => True | _ => False )) then (case s476 s4590 of (Some (s1)) => Some (( 0x042 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s480 s4590)) of Some (s1) => True | _ => False )) then (case s480 s4590 of (Some (s1)) => Some (( 0x043 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s484 s4590)) of Some (s1) => True | _ => False )) then (case s484 s4590 of (Some (s1)) => Some (( 0x044 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s488 s4590)) of Some (s1) => True | _ => False )) then (case s488 s4590 of (Some (s1)) => Some (( 0x001 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s492 s4590)) of Some (s1) => True | _ => False )) then (case s492 s4590 of (Some (s1)) => Some (( 0x002 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s496 s4590)) of Some (s1) => True | _ => False )) then (case s496 s4590 of (Some (s1)) => Some (( 0x003 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s500 s4590)) of Some (s1) => True | _ => False )) then (case s500 s4590 of (Some (s1)) => Some (( 0xC00 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s504 s4590)) of Some (s1) => True | _ => False )) then (case s504 s4590 of (Some (s1)) => Some (( 0xC01 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s508 s4590)) of Some (s1) => True | _ => False )) then (case s508 s4590 of (Some (s1)) => Some (( 0xC02 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s512 s4590)) of Some (s1) => True | _ => False )) then (case s512 s4590 of (Some (s1)) => Some (( 0xC80 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s516 s4590)) of Some (s1) => True | _ => False )) then (case s516 s4590 of (Some (s1)) => Some (( 0xC81 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s520 s4590)) of Some (s1) => True | _ => False )) then (case s520 s4590 of (Some (s1)) => Some (( 0xC82 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s524 s4590)) of Some (s1) => True | _ => False )) then (case s524 s4590 of (Some (s1)) => Some (( 0x100 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s528 s4590)) of Some (s1) => True | _ => False )) then (case s528 s4590 of (Some (s1)) => Some (( 0x102 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s532 s4590)) of Some (s1) => True | _ => False )) then (case s532 s4590 of (Some (s1)) => Some (( 0x103 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s536 s4590)) of Some (s1) => True | _ => False )) then (case s536 s4590 of (Some (s1)) => Some (( 0x104 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s540 s4590)) of Some (s1) => True | _ => False )) then (case s540 s4590 of (Some (s1)) => Some (( 0x105 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s544 s4590)) of Some (s1) => True | _ => False )) then (case s544 s4590 of (Some (s1)) => Some (( 0x106 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s548 s4590)) of Some (s1) => True | _ => False )) then (case s548 s4590 of (Some (s1)) => Some (( 0x140 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s552 s4590)) of Some (s1) => True | _ => False )) then (case s552 s4590 of (Some (s1)) => Some (( 0x141 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s556 s4590)) of Some (s1) => True | _ => False )) then (case s556 s4590 of (Some (s1)) => Some (( 0x142 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s560 s4590)) of Some (s1) => True | _ => False )) then (case s560 s4590 of (Some (s1)) => Some (( 0x143 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s564 s4590)) of Some (s1) => True | _ => False )) then (case s564 s4590 of (Some (s1)) => Some (( 0x144 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s568 s4590)) of Some (s1) => True | _ => False )) then (case s568 s4590 of (Some (s1)) => Some (( 0x180 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s572 s4590)) of Some (s1) => True | _ => False )) then (case s572 s4590 of (Some (s1)) => Some (( 0xF11 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s576 s4590)) of Some (s1) => True | _ => False )) then (case s576 s4590 of (Some (s1)) => Some (( 0xF12 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s580 s4590)) of Some (s1) => True | _ => False )) then (case s580 s4590 of (Some (s1)) => Some (( 0xF13 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s584 s4590)) of Some (s1) => True | _ => False )) then (case s584 s4590 of (Some (s1)) => Some (( 0xF14 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s588 s4590)) of Some (s1) => True | _ => False )) then (case s588 s4590 of (Some (s1)) => Some (( 0x300 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s592 s4590)) of Some (s1) => True | _ => False )) then (case s592 s4590 of (Some (s1)) => Some (( 0x301 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s596 s4590)) of Some (s1) => True | _ => False )) then (case s596 s4590 of (Some (s1)) => Some (( 0x302 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s600 s4590)) of Some (s1) => True | _ => False )) then (case s600 s4590 of (Some (s1)) => Some (( 0x303 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s604 s4590)) of Some (s1) => True | _ => False )) then (case s604 s4590 of (Some (s1)) => Some (( 0x304 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s608 s4590)) of Some (s1) => True | _ => False )) then (case s608 s4590 of (Some (s1)) => Some (( 0x305 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s612 s4590)) of Some (s1) => True | _ => False )) then (case s612 s4590 of (Some (s1)) => Some (( 0x306 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s616 s4590)) of Some (s1) => True | _ => False )) then (case s616 s4590 of (Some (s1)) => Some (( 0x320 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s620 s4590)) of Some (s1) => True | _ => False )) then (case s620 s4590 of (Some (s1)) => Some (( 0x340 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s624 s4590)) of Some (s1) => True | _ => False )) then (case s624 s4590 of (Some (s1)) => Some (( 0x341 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s628 s4590)) of Some (s1) => True | _ => False )) then (case s628 s4590 of (Some (s1)) => Some (( 0x342 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s632 s4590)) of Some (s1) => True | _ => False )) then (case s632 s4590 of (Some (s1)) => Some (( 0x343 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s636 s4590)) of Some (s1) => True | _ => False )) then (case s636 s4590 of (Some (s1)) => Some (( 0x344 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s640 s4590)) of Some (s1) => True | _ => False )) then (case s640 s4590 of (Some (s1)) => Some (( 0x3A0 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s644 s4590)) of Some (s1) => True | _ => False )) then (case s644 s4590 of (Some (s1)) => Some (( 0x3A1 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s648 s4590)) of Some (s1) => True | _ => False )) then (case s648 s4590 of (Some (s1)) => Some (( 0x3A2 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s652 s4590)) of Some (s1) => True | _ => False )) then (case s652 s4590 of (Some (s1)) => Some (( 0x3A3 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s656 s4590)) of Some (s1) => True | _ => False )) then (case s656 s4590 of (Some (s1)) => Some (( 0x3B0 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s660 s4590)) of Some (s1) => True | _ => False )) then (case s660 s4590 of (Some (s1)) => Some (( 0x3B1 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s664 s4590)) of Some (s1) => True | _ => False )) then (case s664 s4590 of (Some (s1)) => Some (( 0x3B2 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s668 s4590)) of Some (s1) => True | _ => False )) then (case s668 s4590 of (Some (s1)) => Some (( 0x3B3 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s672 s4590)) of Some (s1) => True | _ => False )) then (case s672 s4590 of (Some (s1)) => Some (( 0x3B4 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s676 s4590)) of Some (s1) => True | _ => False )) then (case s676 s4590 of (Some (s1)) => Some (( 0x3B5 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s680 s4590)) of Some (s1) => True | _ => False )) then (case s680 s4590 of (Some (s1)) => Some (( 0x3B6 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s684 s4590)) of Some (s1) => True | _ => False )) then (case s684 s4590 of (Some (s1)) => Some (( 0x3B7 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s688 s4590)) of Some (s1) => True | _ => False )) then (case s688 s4590 of (Some (s1)) => Some (( 0x3B8 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s692 s4590)) of Some (s1) => True | _ => False )) then (case s692 s4590 of (Some (s1)) => Some (( 0x3B9 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s696 s4590)) of Some (s1) => True | _ => False )) then (case s696 s4590 of (Some (s1)) => Some (( 0x3BA :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s700 s4590)) of Some (s1) => True | _ => False )) then (case s700 s4590 of (Some (s1)) => Some (( 0x3BB :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s704 s4590)) of Some (s1) => True | _ => False )) then (case s704 s4590 of (Some (s1)) => Some (( 0x3BC :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s708 s4590)) of Some (s1) => True | _ => False )) then (case s708 s4590 of (Some (s1)) => Some (( 0x3BD :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s712 s4590)) of Some (s1) => True | _ => False )) then (case s712 s4590 of (Some (s1)) => Some (( 0x3BE :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s716 s4590)) of Some (s1) => True | _ => False )) then (case s716 s4590 of (Some (s1)) => Some (( 0x3BF :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s720 s4590)) of Some (s1) => True | _ => False )) then (case s720 s4590 of (Some (s1)) => Some (( 0xB00 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s724 s4590)) of Some (s1) => True | _ => False )) then (case s724 s4590 of (Some (s1)) => Some (( 0xB02 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s728 s4590)) of Some (s1) => True | _ => False )) then (case s728 s4590 of (Some (s1)) => Some (( 0xB80 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s732 s4590)) of Some (s1) => True | _ => False )) then (case s732 s4590 of (Some (s1)) => Some (( 0xB82 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s736 s4590)) of Some (s1) => True | _ => False )) then (case s736 s4590 of (Some (s1)) => Some (( 0x7A0 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s740 s4590)) of Some (s1) => True | _ => False )) then (case s740 s4590 of (Some (s1)) => Some (( 0x7A1 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s744 s4590)) of Some (s1) => True | _ => False )) then (case s744 s4590 of (Some (s1)) => Some (( 0x7A2 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s748 s4590)) of Some (s1) => True | _ => False )) then (case s748 s4590 of (Some (s1)) => Some (( 0x7A3 :: 12 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s752 s4590 :: (( 12 Word.word * string))option)) of Some ((reg, s1)) => True | _ => False )) then (case (s752 s4590 :: (( 12 Word.word * string)) option) of (Some ((reg, s1))) => Some (reg, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " definition csr_name :: \(12)Word.word \ string \ where \ csr_name csr = ( csr_name_map_forwards csr )\ for csr :: "(12)Word.word " definition ext_is_CSR_defined :: \(12)Word.word \ Privilege \((register_value),(bool),(exception))monad \ where \ ext_is_CSR_defined b__0 g__326 = ( if (((b__0 = ( 0x000 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) ((haveNExt () )) else if (((b__0 = ( 0x004 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) ((haveNExt () )) else if (((b__0 = ( 0x005 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) ((haveNExt () )) else if (((b__0 = ( 0x040 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) ((haveNExt () )) else if (((b__0 = ( 0x041 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) ((haveNExt () )) else if (((b__0 = ( 0x042 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) ((haveNExt () )) else if (((b__0 = ( 0x043 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) ((haveNExt () )) else if (((b__0 = ( 0x044 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) ((haveNExt () )) else if (((b__0 = ( 0x001 :: 12 Word.word)))) then or_boolM ((haveFExt () )) ((haveDExt () )) else if (((b__0 = ( 0x002 :: 12 Word.word)))) then or_boolM ((haveFExt () )) ((haveDExt () )) else if (((b__0 = ( 0x003 :: 12 Word.word)))) then or_boolM ((haveFExt () )) ((haveDExt () )) else return False )\ for b__0 :: "(12)Word.word " and g__326 :: " Privilege " definition ext_read_CSR :: \(12)Word.word \((register_value),(((32)Word.word)option),(exception))monad \ where \ ext_read_CSR b__0 = ( if (((b__0 = ( 0x000 :: 12 Word.word)))) then read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . return (Some ((get_Ustatus_bits ((lower_sstatus ((lower_mstatus w__0)))) :: 32 Word.word))))) else if (((b__0 = ( 0x004 :: 12 Word.word)))) then read_reg mie_ref \ ((\ (w__1 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__2 :: Minterrupts) . read_reg sideleg_ref \ ((\ (w__3 :: Sinterrupts) . return (Some ((get_Uinterrupts_bits ((lower_sie ((lower_mie w__1 w__2)) w__3)) :: 32 Word.word))))))))) else if (((b__0 = ( 0x005 :: 12 Word.word)))) then (get_utvec () :: ( 32 Word.word) M) \ ((\ (w__4 :: 32 Word.word) . return (Some w__4))) else if (((b__0 = ( 0x040 :: 12 Word.word)))) then (read_reg uscratch_ref :: ( 32 Word.word) M) \ ((\ (w__5 :: 32 Word.word) . return (Some w__5))) else if (((b__0 = ( 0x041 :: 12 Word.word)))) then (get_xret_target User :: ( 32 Word.word) M) \ ((\ (w__6 :: 32 Word.word) . (pc_alignment_mask () :: ( 32 Word.word) M) \ ((\ (w__7 :: 32 Word.word) . return (Some ((and_vec w__6 w__7 :: 32 Word.word))))))) else if (((b__0 = ( 0x042 :: 12 Word.word)))) then read_reg ucause_ref \ ((\ (w__8 :: Mcause) . return (Some ((get_Mcause_bits w__8 :: 32 Word.word))))) else if (((b__0 = ( 0x043 :: 12 Word.word)))) then (read_reg utval_ref :: ( 32 Word.word) M) \ ((\ (w__9 :: 32 Word.word) . return (Some w__9))) else if (((b__0 = ( 0x044 :: 12 Word.word)))) then read_reg mip_ref \ ((\ (w__10 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__11 :: Minterrupts) . read_reg sideleg_ref \ ((\ (w__12 :: Sinterrupts) . return (Some ((get_Uinterrupts_bits ((lower_sip ((lower_mip w__10 w__11)) w__12)) :: 32 Word.word))))))))) else if (((b__0 = ( 0x001 :: 12 Word.word)))) then read_reg fcsr_ref \ ((\ (w__13 :: Fcsr) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Fcsr_FFLAGS w__13 :: 5 Word.word)) :: 32 Word.word))))) else if (((b__0 = ( 0x002 :: 12 Word.word)))) then read_reg fcsr_ref \ ((\ (w__14 :: Fcsr) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Fcsr_FRM w__14 :: 3 Word.word)) :: 32 Word.word))))) else if (((b__0 = ( 0x003 :: 12 Word.word)))) then read_reg fcsr_ref \ ((\ (w__15 :: Fcsr) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Fcsr_bits w__15 :: 32 Word.word)) :: 32 Word.word))))) else return None )\ for b__0 :: "(12)Word.word " definition ext_write_CSR :: \(12)Word.word \(32)Word.word \((register_value),(((32)Word.word)option),(exception))monad \ where \ ext_write_CSR b__0 value1 = ( if (((b__0 = ( 0x000 :: 12 Word.word)))) then read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . legalize_ustatus w__0 value1 \ ((\ (w__1 :: Mstatus) . (write_reg mstatus_ref w__1 \ read_reg mstatus_ref) \ ((\ (w__2 :: Mstatus) . return (Some ((get_Mstatus_bits w__2 :: 32 Word.word))))))))) else if (((b__0 = ( 0x004 :: 12 Word.word)))) then read_reg mie_ref \ ((\ (w__3 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__4 :: Minterrupts) . read_reg sideleg_ref \ ((\ (w__5 :: Sinterrupts) . (let sie = (legalize_uie ((lower_mie w__3 w__4)) w__5 value1) in read_reg mie_ref \ ((\ (w__6 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__7 :: Minterrupts) . lift_sie w__6 w__7 sie \ ((\ (w__8 :: Minterrupts) . (write_reg mie_ref w__8 \ read_reg mie_ref) \ ((\ (w__9 :: Minterrupts) . return (Some ((get_Minterrupts_bits w__9 :: 32 Word.word)))))))))))))))))) else if (((b__0 = ( 0x005 :: 12 Word.word)))) then (set_utvec value1 :: ( 32 Word.word) M) \ ((\ (w__10 :: 32 Word.word) . return (Some w__10))) else if (((b__0 = ( 0x040 :: 12 Word.word)))) then (write_reg uscratch_ref value1 \ (read_reg uscratch_ref :: ( 32 Word.word) M)) \ ((\ (w__11 :: 32 Word.word) . return (Some w__11))) else if (((b__0 = ( 0x041 :: 12 Word.word)))) then (set_xret_target User value1 :: ( 32 Word.word) M) \ ((\ (w__12 :: 32 Word.word) . return (Some w__12))) else if (((b__0 = ( 0x042 :: 12 Word.word)))) then (set_Mcause_bits ucause_ref value1 \ read_reg ucause_ref) \ ((\ (w__13 :: Mcause) . return (Some ((get_Mcause_bits w__13 :: 32 Word.word))))) else if (((b__0 = ( 0x043 :: 12 Word.word)))) then (write_reg utval_ref value1 \ (read_reg utval_ref :: ( 32 Word.word) M)) \ ((\ (w__14 :: 32 Word.word) . return (Some w__14))) else if (((b__0 = ( 0x044 :: 12 Word.word)))) then read_reg mip_ref \ ((\ (w__15 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__16 :: Minterrupts) . read_reg sideleg_ref \ ((\ (w__17 :: Sinterrupts) . (let sip = (legalize_uip ((lower_mip w__15 w__16)) w__17 value1) in read_reg mip_ref \ ((\ (w__18 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__19 :: Minterrupts) . lift_sip w__18 w__19 sip \ ((\ (w__20 :: Minterrupts) . (write_reg mip_ref w__20 \ read_reg mip_ref) \ ((\ (w__21 :: Minterrupts) . return (Some ((get_Minterrupts_bits w__21 :: 32 Word.word)))))))))))))))))) else if (((b__0 = ( 0x001 :: 12 Word.word)))) then read_reg fcsr_ref \ ((\ (w__22 :: Fcsr) . (ext_write_fcsr ((get_Fcsr_FRM w__22 :: 3 Word.word)) ((subrange_vec_dec value1 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) \ read_reg fcsr_ref) \ ((\ (w__23 :: Fcsr) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Fcsr_FFLAGS w__23 :: 5 Word.word)) :: 32 Word.word))))))) else if (((b__0 = ( 0x002 :: 12 Word.word)))) then read_reg fcsr_ref \ ((\ (w__24 :: Fcsr) . (ext_write_fcsr ((subrange_vec_dec value1 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) ((get_Fcsr_FFLAGS w__24 :: 5 Word.word)) \ read_reg fcsr_ref) \ ((\ (w__25 :: Fcsr) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Fcsr_FRM w__25 :: 3 Word.word)) :: 32 Word.word))))))) else if (((b__0 = ( 0x003 :: 12 Word.word)))) then (ext_write_fcsr ((subrange_vec_dec value1 (( 7 :: int)::ii) (( 5 :: int)::ii) :: 3 Word.word)) ((subrange_vec_dec value1 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) \ read_reg fcsr_ref) \ ((\ (w__26 :: Fcsr) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Fcsr_bits w__26 :: 32 Word.word)) :: 32 Word.word))))) else return None )\ for b__0 :: "(12)Word.word " and value1 :: "(32)Word.word " \ \\val csrAccess : mword ty12 -> mword ty2\\ definition csrAccess :: \(12)Word.word \(2)Word.word \ where \ csrAccess csr = ( (subrange_vec_dec csr (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word))\ for csr :: "(12)Word.word " \ \\val csrPriv : mword ty12 -> mword ty2\\ definition csrPriv :: \(12)Word.word \(2)Word.word \ where \ csrPriv csr = ( (subrange_vec_dec csr (( 9 :: int)::ii) (( 8 :: int)::ii) :: 2 Word.word))\ for csr :: "(12)Word.word " \ \\val is_CSR_defined : mword ty12 -> Privilege -> M bool\\ definition is_CSR_defined :: \(12)Word.word \ Privilege \((register_value),(bool),(exception))monad \ where \ is_CSR_defined (csr :: csreg) (p :: Privilege) = ( (let b__0 = csr in if (((b__0 = ( 0xF11 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0xF12 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0xF13 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0xF14 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x300 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x301 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x302 :: 12 Word.word)))) then and_boolM (return (((p = Machine)))) (or_boolM ((haveSupMode () )) ((haveNExt () ))) else if (((b__0 = ( 0x303 :: 12 Word.word)))) then and_boolM (return (((p = Machine)))) (or_boolM ((haveSupMode () )) ((haveNExt () ))) else if (((b__0 = ( 0x304 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x305 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x306 :: 12 Word.word)))) then and_boolM (return (((p = Machine)))) ((haveUsrMode () )) else if (((b__0 = ( 0x310 :: 12 Word.word)))) then return ((((((p = Machine))) \ ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))))) else if (((b__0 = ( 0x320 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x340 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x341 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x342 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x343 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x344 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3A0 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3A1 :: 12 Word.word)))) then return ((((((p = Machine))) \ ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))))) else if (((b__0 = ( 0x3A2 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3A3 :: 12 Word.word)))) then return ((((((p = Machine))) \ ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))))) else if (((b__0 = ( 0x3B0 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B1 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B2 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B3 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B4 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B5 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B6 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B7 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B8 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3B9 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3BA :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3BB :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3BC :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3BD :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3BE :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x3BF :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0xB00 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0xB02 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0xB80 :: 12 Word.word)))) then return ((((((p = Machine))) \ ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))))) else if (((b__0 = ( 0xB82 :: 12 Word.word)))) then return ((((((p = Machine))) \ ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))))) else if (((b__0 = ( 0x7A0 :: 12 Word.word)))) then return (((p = Machine))) else if (((b__0 = ( 0x100 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x102 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (and_boolM ((haveNExt () )) (return ((((((p = Machine))) \ (((p = Supervisor)))))))) else if (((b__0 = ( 0x103 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (and_boolM ((haveNExt () )) (return ((((((p = Machine))) \ (((p = Supervisor)))))))) else if (((b__0 = ( 0x104 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x105 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x106 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x140 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x141 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x142 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x143 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x144 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0x180 :: 12 Word.word)))) then and_boolM ((haveSupMode () )) (return ((((((p = Machine))) \ (((p = Supervisor))))))) else if (((b__0 = ( 0xC00 :: 12 Word.word)))) then haveUsrMode () else if (((b__0 = ( 0xC01 :: 12 Word.word)))) then haveUsrMode () else if (((b__0 = ( 0xC02 :: 12 Word.word)))) then haveUsrMode () else if (((b__0 = ( 0xC80 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) else if (((b__0 = ( 0xC81 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) else if (((b__0 = ( 0xC82 :: 12 Word.word)))) then and_boolM ((haveUsrMode () )) (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) else ext_is_CSR_defined csr p))\ for csr :: "(12)Word.word " and p :: " Privilege " \ \\val check_CSR_access : mword ty2 -> mword ty2 -> Privilege -> bool -> bool\\ definition check_CSR_access :: \(2)Word.word \(2)Word.word \ Privilege \ bool \ bool \ where \ check_CSR_access csrrw csrpr p isWrite = ( (((\ ((((((isWrite = True))) \ (((csrrw = ( 0b11 :: 2 Word.word))))))))) \ ((zopz0zKzJ_u ((privLevel_to_bits p :: 2 Word.word)) csrpr))))\ for csrrw :: "(2)Word.word " and csrpr :: "(2)Word.word " and p :: " Privilege " and isWrite :: " bool " \ \\val check_TVM_SATP : mword ty12 -> Privilege -> M bool\\ definition check_TVM_SATP :: \(12)Word.word \ Privilege \((register_value),(bool),(exception))monad \ where \ check_TVM_SATP (csr :: csreg) (p :: Privilege) = ( and_boolM (return (((csr = ( 0x180 :: 12 Word.word))))) (and_boolM (return (((p = Supervisor)))) (read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . return (((((get_Mstatus_TVM w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))) \ ((\ (w__2 :: bool) . return ((\ w__2)))))\ for csr :: "(12)Word.word " and p :: " Privilege " \ \\val check_Counteren : mword ty12 -> Privilege -> M bool\\ definition check_Counteren :: \(12)Word.word \ Privilege \((register_value),(bool),(exception))monad \ where \ check_Counteren (csr :: csreg) (p :: Privilege) = ( (case (csr, p) of (b__0, Supervisor) => if (((b__0 = ( 0xC00 :: 12 Word.word)))) then read_reg mcounteren_ref \ ((\ (w__0 :: Counteren) . return (((((get_Counteren_CY w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))) else if (((b__0 = ( 0xC01 :: 12 Word.word)))) then read_reg mcounteren_ref \ ((\ (w__1 :: Counteren) . return (((((get_Counteren_TM w__1 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))) else if (((b__0 = ( 0xC02 :: 12 Word.word)))) then read_reg mcounteren_ref \ ((\ (w__2 :: Counteren) . return (((((get_Counteren_IR w__2 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))) else return ((case (b__0, Supervisor) of (_, _) => if (((((zopz0zIzJ_u ( 0xC03 :: 12 Word.word) csr)) \ ((zopz0zIzJ_u csr ( 0xC1F :: 12 Word.word)))))) then False else True )) | (b__3, User) => if (((b__3 = ( 0xC00 :: 12 Word.word)))) then and_boolM (read_reg mcounteren_ref \ ((\ (w__6 :: Counteren) . return (((((get_Counteren_CY w__6 :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))) (or_boolM (haveSupMode () \ ((\ (w__7 :: bool) . return ((\ w__7))))) (read_reg scounteren_ref \ ((\ (w__8 :: Counteren) . return (((((get_Counteren_CY w__8 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))) else if (((b__3 = ( 0xC01 :: 12 Word.word)))) then and_boolM (read_reg mcounteren_ref \ ((\ (w__11 :: Counteren) . return (((((get_Counteren_TM w__11 :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))) (or_boolM (haveSupMode () \ ((\ (w__12 :: bool) . return ((\ w__12))))) (read_reg scounteren_ref \ ((\ (w__13 :: Counteren) . return (((((get_Counteren_TM w__13 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))) else if (((b__3 = ( 0xC02 :: 12 Word.word)))) then and_boolM (read_reg mcounteren_ref \ ((\ (w__16 :: Counteren) . return (((((get_Counteren_IR w__16 :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))) (or_boolM (haveSupMode () \ ((\ (w__17 :: bool) . return ((\ w__17))))) (read_reg scounteren_ref \ ((\ (w__18 :: Counteren) . return (((((get_Counteren_IR w__18 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))) else return ((case (b__3, User) of (_, _) => if (((((zopz0zIzJ_u ( 0xC03 :: 12 Word.word) csr)) \ ((zopz0zIzJ_u csr ( 0xC1F :: 12 Word.word)))))) then False else True )) | (_, _) => return (if (((((zopz0zIzJ_u ( 0xC03 :: 12 Word.word) csr)) \ ((zopz0zIzJ_u csr ( 0xC1F :: 12 Word.word)))))) then False else True) ))\ for csr :: "(12)Word.word " and p :: " Privilege " \ \\val check_CSR : mword ty12 -> Privilege -> bool -> M bool\\ definition check_CSR :: \(12)Word.word \ Privilege \ bool \((register_value),(bool),(exception))monad \ where \ check_CSR (csr :: csreg) (p :: Privilege) (isWrite :: bool) = ( and_boolM ((is_CSR_defined csr p)) (and_boolM (return ((check_CSR_access ((csrAccess csr :: 2 Word.word)) ((csrPriv csr :: 2 Word.word)) p isWrite))) (and_boolM ((check_TVM_SATP csr p)) ((check_Counteren csr p)))))\ for csr :: "(12)Word.word " and p :: " Privilege " and isWrite :: " bool " \ \\val exception_delegatee : ExceptionType -> Privilege -> M Privilege\\ definition exception_delegatee :: \ ExceptionType \ Privilege \((register_value),(Privilege),(exception))monad \ where \ exception_delegatee (e :: ExceptionType) (p :: Privilege) = ( (let idx = (num_of_ExceptionType e) in read_reg medeleg_ref \ ((\ (w__0 :: Medeleg) . (let super = (bit_to_bool ((access_vec_dec ((get_Medeleg_bits w__0 :: 32 Word.word)) idx))) in haveSupMode () \ ((\ (w__1 :: bool) . (if w__1 then and_boolM (return super) (and_boolM ((haveNExt () )) (read_reg sedeleg_ref \ ((\ (w__3 :: Sedeleg) . return ((bit_to_bool ((access_vec_dec ((get_Sedeleg_bits w__3 :: 32 Word.word)) idx)))))))) else and_boolM (return super) ((haveNExt () ))) \ ((\ user . and_boolM ((haveUsrMode () )) (return user) \ ((\ w__9 . (if w__9 then return User else and_boolM ((haveSupMode () )) (return super) \ ((\ w__11 . return (if w__11 then Supervisor else Machine)))) \ ((\ deleg . return (if ((zopz0zI_u ((privLevel_to_bits deleg :: 2 Word.word)) ((privLevel_to_bits p :: 2 Word.word)))) then p else deleg))))))))))))))\ for e :: " ExceptionType " and p :: " Privilege " \ \\val findPendingInterrupt : mword ty32 -> maybe InterruptType\\ definition findPendingInterrupt :: \(32)Word.word \(InterruptType)option \ where \ findPendingInterrupt ip = ( (let ip = (Mk_Minterrupts ip) in if (((((get_Minterrupts_MEI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_M_External else if (((((get_Minterrupts_MSI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_M_Software else if (((((get_Minterrupts_MTI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_M_Timer else if (((((get_Minterrupts_SEI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_S_External else if (((((get_Minterrupts_SSI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_S_Software else if (((((get_Minterrupts_STI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_S_Timer else if (((((get_Minterrupts_UEI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_U_External else if (((((get_Minterrupts_USI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_U_Software else if (((((get_Minterrupts_UTI ip :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Some I_U_Timer else None))\ for ip :: "(32)Word.word " \ \\val processPending : Minterrupts -> Minterrupts -> mword ty32 -> bool -> interrupt_set\\ definition processPending :: \ Minterrupts \ Minterrupts \(32)Word.word \ bool \ interrupt_set \ where \ processPending (xip :: Minterrupts) (xie :: Minterrupts) (xideleg :: xlenbits) (priv_enabled :: bool) = ( (let effective_pend = ((and_vec ((get_Minterrupts_bits xip :: 32 Word.word)) ((and_vec ((get_Minterrupts_bits xie :: 32 Word.word)) ((not_vec xideleg :: 32 Word.word)) :: 32 Word.word)) :: 32 Word.word)) in (let effective_delg = ((and_vec ((get_Minterrupts_bits xip :: 32 Word.word)) xideleg :: 32 Word.word)) in if (((priv_enabled \ (((effective_pend \ ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word)))))))) then Ints_Pending effective_pend else if (((effective_delg \ ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))))) then Ints_Delegated effective_delg else Ints_Empty () )))\ for xip :: " Minterrupts " and xie :: " Minterrupts " and xideleg :: "(32)Word.word " and priv_enabled :: " bool " \ \\val getPendingSet : Privilege -> M (maybe ((mword ty32 * Privilege)))\\ definition getPendingSet :: \ Privilege \((register_value),(((32)Word.word*Privilege)option),(exception))monad \ where \ getPendingSet priv = ( haveUsrMode () \ ((\ (w__0 :: bool) . (assert_exp w__0 (''no user mode: M/U or M/S/U system required'') \ read_reg mip_ref) \ ((\ (w__1 :: Minterrupts) . read_reg mie_ref \ ((\ (w__2 :: Minterrupts) . (let effective_pending = ((and_vec ((get_Minterrupts_bits w__1 :: 32 Word.word)) ((get_Minterrupts_bits w__2 :: 32 Word.word)) :: 32 Word.word)) in if (((effective_pending = ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))))) then return None else or_boolM (return (((priv \ Machine)))) (and_boolM (return (((priv = Machine)))) (read_reg mstatus_ref \ ((\ (w__3 :: Mstatus) . return (((((get_Mstatus_MIE w__3 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))) \ ((\ mIE . and_boolM ((haveSupMode () )) (or_boolM (return (((priv = User)))) (and_boolM (return (((priv = Supervisor)))) (read_reg mstatus_ref \ ((\ (w__6 :: Mstatus) . return (((((get_Mstatus_SIE w__6 :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))))) \ ((\ sIE . and_boolM ((haveNExt () )) (and_boolM (return (((priv = User)))) (read_reg mstatus_ref \ ((\ (w__10 :: Mstatus) . return (((((get_Mstatus_UIE w__10 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))))))) \ ((\ uIE . read_reg mip_ref \ ((\ (w__12 :: Minterrupts) . read_reg mie_ref \ ((\ (w__13 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__14 :: Minterrupts) . (case ((processPending w__12 w__13 ((get_Minterrupts_bits w__14 :: 32 Word.word)) mIE)) of Ints_Empty (_) => return None | Ints_Pending (p) => (let r = (p, Machine) in return (Some r)) | Ints_Delegated (d) => haveSupMode () \ ((\ (w__15 :: bool) . if ((\ w__15)) then return (if uIE then (let r = (d, User) in Some r) else None) else read_reg mie_ref \ ((\ (w__16 :: Minterrupts) . read_reg sideleg_ref \ ((\ (w__17 :: Sinterrupts) . return ((case ((processPending ((Mk_Minterrupts d)) w__16 ((get_Sinterrupts_bits w__17 :: 32 Word.word)) sIE)) of Ints_Empty (_) => None | Ints_Pending (p) => (let r = (p, Supervisor) in Some r) | Ints_Delegated (d) => if uIE then (let r = (d, User) in Some r) else None )))))))) )))))))))))))))))))))\ for priv :: " Privilege " \ \\val dispatchInterrupt : Privilege -> M (maybe ((InterruptType * Privilege)))\\ definition dispatchInterrupt :: \ Privilege \((register_value),((InterruptType*Privilege)option),(exception))monad \ where \ dispatchInterrupt priv = ( or_boolM (haveUsrMode () \ ((\ (w__0 :: bool) . return ((\ w__0))))) (and_boolM (haveSupMode () \ ((\ (w__1 :: bool) . return ((\ w__1))))) (haveNExt () \ ((\ (w__2 :: bool) . return ((\ w__2)))))) \ ((\ (w__4 :: bool) . if w__4 then (assert_exp (((priv = Machine))) (''invalid current privilege'') \ read_reg mip_ref) \ ((\ (w__5 :: Minterrupts) . read_reg mie_ref \ ((\ (w__6 :: Minterrupts) . (let enabled_pending = ((and_vec ((get_Minterrupts_bits w__5 :: 32 Word.word)) ((get_Minterrupts_bits w__6 :: 32 Word.word)) :: 32 Word.word)) in return ((case ((findPendingInterrupt enabled_pending)) of Some (i) => (let r = (i, Machine) in Some r) | None => None ))))))) else (getPendingSet priv :: ( (( 32 Word.word * Privilege))option) M) \ ((\ (w__7 :: (( 32 Word.word * Privilege))option) . return ((case w__7 of None => None | Some ((ip, p)) => (case ((findPendingInterrupt ip)) of None => None | Some (i) => (let r = (i, p) in Some r) ) )))))))\ for priv :: " Privilege " \ \\val tval : maybe (mword ty32) -> mword ty32\\ fun tval :: \((32)Word.word)option \(32)Word.word \ where \ tval (Some (e)) = ( e )\ for e :: "(32)Word.word " |\ tval None = ( (EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))\ \ \\val rvfi_trap : unit -> unit\\ definition rvfi_trap :: \ unit \ unit \ where \ rvfi_trap _ = ( () )\ \ \\val trap_handler : Privilege -> bool -> mword ty8 -> mword ty32 -> maybe (mword ty32) -> maybe unit -> M (mword ty32)\\ definition trap_handler :: \ Privilege \ bool \(8)Word.word \(32)Word.word \(xlenbits)option \(ext_exception)option \((register_value),((32)Word.word),(exception))monad \ where \ trap_handler (del_priv :: Privilege) (intr :: bool) (c :: exc_code) (pc :: xlenbits) (info :: xlenbits option) (ext :: ext_exception option) = ( (let (_ :: unit) = (rvfi_trap () ) in (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''handling '') (((@) (if intr then (''int#'') else (''exc#'')) (((@) ((string_of_bits c)) (((@) ('' at priv '') (((@) ((privLevel_to_str del_priv)) (((@) ('' with tval '') ((string_of_bits ((tval info :: 32 Word.word)))))))))))))))) else () ) in (let (_ :: unit) = (cancel_reservation () ) in (case del_priv of Machine => ((set_Mcause_IsInterrupt mcause_ref ((bool_to_bits intr :: 1 Word.word)) \ set_Mcause_Cause mcause_ref ((EXTZ (( 31 :: int)::ii) c :: 31 Word.word))) \ read_reg mstatus_ref) \ ((\ (w__0 :: Mstatus) . ((set_Mstatus_MPIE mstatus_ref ((get_Mstatus_MIE w__0 :: 1 Word.word)) \ set_Mstatus_MIE mstatus_ref ( 0b0 :: 1 Word.word)) \ read_reg cur_privilege_ref) \ ((\ (w__1 :: Privilege) . (((set_Mstatus_MPP mstatus_ref ((privLevel_to_bits w__1 :: 2 Word.word)) \ write_reg mtval_ref ((tval info :: 32 Word.word))) \ write_reg mepc_ref pc) \ write_reg cur_privilege_ref del_priv) \ ((let (_ :: unit) = (handle_trap_extension del_priv pc ext) in ((if ((get_config_print_reg () )) then read_reg mstatus_ref \ ((\ (w__2 :: Mstatus) . return ((print_dbg (((@) (''CSR mstatus <- '') ((string_of_bits ((get_Mstatus_bits w__2 :: 32 Word.word)))))))))) else return () ) \ read_reg mcause_ref) \ ((\ (w__3 :: Mcause) . (prepare_trap_vector del_priv w__3 :: ( 32 Word.word) M))))))))) | Supervisor => haveSupMode () \ ((\ (w__5 :: bool) . (((assert_exp w__5 (''no supervisor mode present for delegation'') \ set_Mcause_IsInterrupt scause_ref ((bool_to_bits intr :: 1 Word.word))) \ set_Mcause_Cause scause_ref ((EXTZ (( 31 :: int)::ii) c :: 31 Word.word))) \ read_reg mstatus_ref) \ ((\ (w__6 :: Mstatus) . ((set_Mstatus_SPIE mstatus_ref ((get_Mstatus_SIE w__6 :: 1 Word.word)) \ set_Mstatus_SIE mstatus_ref ( 0b0 :: 1 Word.word)) \ read_reg cur_privilege_ref) \ ((\ (w__7 :: Privilege) . (case w__7 of User => return ( 0b0 :: 1 Word.word) | Supervisor => return ( 0b1 :: 1 Word.word) | Machine => (internal_error (''invalid privilege for s-mode trap'') :: ( 1 Word.word) M) ) \ ((\ (w__9 :: 1 Word.word) . (((set_Mstatus_SPP mstatus_ref w__9 \ write_reg stval_ref ((tval info :: 32 Word.word))) \ write_reg sepc_ref pc) \ write_reg cur_privilege_ref del_priv) \ ((let (_ :: unit) = (handle_trap_extension del_priv pc ext) in ((if ((get_config_print_reg () )) then read_reg mstatus_ref \ ((\ (w__10 :: Mstatus) . return ((print_dbg (((@) (''CSR mstatus <- '') ((string_of_bits ((get_Mstatus_bits w__10 :: 32 Word.word)))))))))) else return () ) \ read_reg scause_ref) \ ((\ (w__11 :: Mcause) . (prepare_trap_vector del_priv w__11 :: ( 32 Word.word) M))))))))))))) | User => haveUsrMode () \ ((\ (w__13 :: bool) . (((assert_exp w__13 (''no user mode present for delegation'') \ set_Mcause_IsInterrupt ucause_ref ((bool_to_bits intr :: 1 Word.word))) \ set_Mcause_Cause ucause_ref ((EXTZ (( 31 :: int)::ii) c :: 31 Word.word))) \ read_reg mstatus_ref) \ ((\ (w__14 :: Mstatus) . ((((set_Mstatus_UPIE mstatus_ref ((get_Mstatus_UIE w__14 :: 1 Word.word)) \ set_Mstatus_UIE mstatus_ref ( 0b0 :: 1 Word.word)) \ write_reg utval_ref ((tval info :: 32 Word.word))) \ write_reg uepc_ref pc) \ write_reg cur_privilege_ref del_priv) \ ((let (_ :: unit) = (handle_trap_extension del_priv pc ext) in ((if ((get_config_print_reg () )) then read_reg mstatus_ref \ ((\ (w__15 :: Mstatus) . return ((print_dbg (((@) (''CSR mstatus <- '') ((string_of_bits ((get_Mstatus_bits w__15 :: 32 Word.word)))))))))) else return () ) \ read_reg ucause_ref) \ ((\ (w__16 :: Mcause) . (prepare_trap_vector del_priv w__16 :: ( 32 Word.word) M))))))))) )))))\ for del_priv :: " Privilege " and intr :: " bool " and c :: "(8)Word.word " and pc :: "(32)Word.word " and info :: "(xlenbits)option " and ext :: "(ext_exception)option " \ \\val exception_handler : Privilege -> ctl_result -> mword ty32 -> M (mword ty32)\\ definition exception_handler :: \ Privilege \ ctl_result \(32)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ exception_handler (cur_priv :: Privilege) (ctl :: ctl_result) (pc :: xlenbits) = ( (case (cur_priv, ctl) of (_, CTL_TRAP (e)) => exception_delegatee(sync_exception_trap e) cur_priv \ ((\ del_priv . (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''trapping from '') (((@) ((privLevel_to_str cur_priv)) (((@) ('' to '') (((@) ((privLevel_to_str del_priv)) (((@) ('' to handle '') ((exceptionType_to_str(sync_exception_trap e))))))))))))) else () ) in (trap_handler del_priv False ((exceptionType_to_bits(sync_exception_trap e) :: 8 Word.word)) pc(sync_exception_excinfo e)(sync_exception_ext_exception e) :: ( 32 Word.word) M)))) | (_, CTL_MRET (_)) => read_reg cur_privilege_ref \ ((\ prev_priv . read_reg mstatus_ref \ ((\ (w__1 :: Mstatus) . ((set_Mstatus_MIE mstatus_ref ((get_Mstatus_MPIE w__1 :: 1 Word.word)) \ set_Mstatus_MPIE mstatus_ref ( 0b1 :: 1 Word.word)) \ read_reg mstatus_ref) \ ((\ (w__2 :: Mstatus) . privLevel_of_bits ((get_Mstatus_MPP w__2 :: 2 Word.word)) \ ((\ (w__3 :: Privilege) . (write_reg cur_privilege_ref w__3 \ haveUsrMode () ) \ ((\ (w__4 :: bool) . (set_Mstatus_MPP mstatus_ref ((privLevel_to_bits (if w__4 then User else Machine) :: 2 Word.word)) \ read_reg cur_privilege_ref) \ ((\ (w__5 :: Privilege) . (((if (((w__5 \ Machine))) then set_Mstatus_MPRV mstatus_ref ( 0b0 :: 1 Word.word) else return () ) \ (if ((get_config_print_reg () )) then read_reg mstatus_ref \ ((\ (w__6 :: Mstatus) . return ((print_dbg (((@) (''CSR mstatus <- '') ((string_of_bits ((get_Mstatus_bits w__6 :: 32 Word.word)))))))))) else return () )) \ (if ((get_config_print_platform () )) then read_reg cur_privilege_ref \ ((\ (w__7 :: Privilege) . return ((print_dbg (((@) (''ret-ing from '') (((@) ((privLevel_to_str prev_priv)) (((@) ('' to '') ((privLevel_to_str w__7)))))))))))) else return () )) \ ((let (_ :: unit) = (cancel_reservation () ) in (prepare_xret_target Machine :: ( 32 Word.word) M) \ ((\ (w__8 :: 32 Word.word) . (pc_alignment_mask () :: ( 32 Word.word) M) \ ((\ (w__9 :: 32 Word.word) . return ((and_vec w__8 w__9 :: 32 Word.word)))))))))))))))))))) | (_, CTL_SRET (_)) => read_reg cur_privilege_ref \ ((\ prev_priv . read_reg mstatus_ref \ ((\ (w__10 :: Mstatus) . ((set_Mstatus_SIE mstatus_ref ((get_Mstatus_SPIE w__10 :: 1 Word.word)) \ set_Mstatus_SPIE mstatus_ref ( 0b1 :: 1 Word.word)) \ read_reg mstatus_ref) \ ((\ (w__11 :: Mstatus) . ((write_reg cur_privilege_ref (if (((((get_Mstatus_SPP w__11 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then Supervisor else User) \ set_Mstatus_SPP mstatus_ref ( 0b0 :: 1 Word.word)) \ read_reg cur_privilege_ref) \ ((\ (w__12 :: Privilege) . (((if (((w__12 \ Machine))) then set_Mstatus_MPRV mstatus_ref ( 0b0 :: 1 Word.word) else return () ) \ (if ((get_config_print_reg () )) then read_reg mstatus_ref \ ((\ (w__13 :: Mstatus) . return ((print_dbg (((@) (''CSR mstatus <- '') ((string_of_bits ((get_Mstatus_bits w__13 :: 32 Word.word)))))))))) else return () )) \ (if ((get_config_print_platform () )) then read_reg cur_privilege_ref \ ((\ (w__14 :: Privilege) . return ((print_dbg (((@) (''ret-ing from '') (((@) ((privLevel_to_str prev_priv)) (((@) ('' to '') ((privLevel_to_str w__14)))))))))))) else return () )) \ ((let (_ :: unit) = (cancel_reservation () ) in (prepare_xret_target Supervisor :: ( 32 Word.word) M) \ ((\ (w__15 :: 32 Word.word) . (pc_alignment_mask () :: ( 32 Word.word) M) \ ((\ (w__16 :: 32 Word.word) . return ((and_vec w__15 w__16 :: 32 Word.word)))))))))))))))) | (_, CTL_URET (_)) => read_reg cur_privilege_ref \ ((\ prev_priv . read_reg mstatus_ref \ ((\ (w__17 :: Mstatus) . ((((set_Mstatus_UIE mstatus_ref ((get_Mstatus_UPIE w__17 :: 1 Word.word)) \ set_Mstatus_UPIE mstatus_ref ( 0b1 :: 1 Word.word)) \ write_reg cur_privilege_ref User) \ (if ((get_config_print_reg () )) then read_reg mstatus_ref \ ((\ (w__18 :: Mstatus) . return ((print_dbg (((@) (''CSR mstatus <- '') ((string_of_bits ((get_Mstatus_bits w__18 :: 32 Word.word)))))))))) else return () )) \ (if ((get_config_print_platform () )) then read_reg cur_privilege_ref \ ((\ (w__19 :: Privilege) . return ((print_dbg (((@) (''ret-ing from '') (((@) ((privLevel_to_str prev_priv)) (((@) ('' to '') ((privLevel_to_str w__19)))))))))))) else return () )) \ ((let (_ :: unit) = (cancel_reservation () ) in (prepare_xret_target User :: ( 32 Word.word) M) \ ((\ (w__20 :: 32 Word.word) . (pc_alignment_mask () :: ( 32 Word.word) M) \ ((\ (w__21 :: 32 Word.word) . return ((and_vec w__20 w__21 :: 32 Word.word)))))))))))) ))\ for cur_priv :: " Privilege " and ctl :: " ctl_result " and pc :: "(32)Word.word " \ \\val handle_mem_exception : mword ty32 -> ExceptionType -> M unit\\ definition handle_mem_exception :: \(32)Word.word \ ExceptionType \((register_value),(unit),(exception))monad \ where \ handle_mem_exception (addr :: xlenbits) (e :: ExceptionType) = ( (let (t :: sync_exception) = ((| sync_exception_trap = e, sync_exception_excinfo = (Some addr), sync_exception_ext_exception = None |)) in read_reg cur_privilege_ref \ ((\ (w__0 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (exception_handler w__0 (CTL_TRAP t) w__1 :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . set_next_pc w__2))))))))\ for addr :: "(32)Word.word " and e :: " ExceptionType " \ \\val handle_exception : ExceptionType -> M unit\\ definition handle_exception :: \ ExceptionType \((register_value),(unit),(exception))monad \ where \ handle_exception e = ( (let (t :: sync_exception) = ((| sync_exception_trap = e, sync_exception_excinfo = None, sync_exception_ext_exception = None |)) in read_reg cur_privilege_ref \ ((\ (w__0 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (exception_handler w__0 (CTL_TRAP t) w__1 :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . set_next_pc w__2))))))))\ for e :: " ExceptionType " \ \\val handle_interrupt : InterruptType -> Privilege -> M unit\\ definition handle_interrupt :: \ InterruptType \ Privilege \((register_value),(unit),(exception))monad \ where \ handle_interrupt (i :: InterruptType) (del_priv :: Privilege) = ( (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (trap_handler del_priv True ((interruptType_to_bits i :: 8 Word.word)) w__0 None None :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . set_next_pc w__1)))))\ for i :: " InterruptType " and del_priv :: " Privilege " \ \\val init_sys : unit -> M unit\\ definition init_sys :: \ unit \((register_value),(unit),(exception))monad \ where \ init_sys _ = ( (((((((((((write_reg cur_privilege_ref Machine \ write_reg mhartid_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ set_Misa_MXL misa_ref ((arch_to_bits RV32 :: 2 Word.word))) \ set_Misa_A misa_ref ( 0b1 :: 1 Word.word)) \ set_Misa_C misa_ref ((bool_to_bits ((sys_enable_rvc () )) :: 1 Word.word))) \ set_Misa_I misa_ref ( 0b1 :: 1 Word.word)) \ set_Misa_M misa_ref ( 0b1 :: 1 Word.word)) \ set_Misa_U misa_ref ( 0b1 :: 1 Word.word)) \ set_Misa_S misa_ref ( 0b1 :: 1 Word.word)) \ set_Misa_F misa_ref ((bool_to_bits ((sys_enable_fdext () )) :: 1 Word.word))) \ set_Misa_D misa_ref ( 0b0 :: 1 Word.word)) \ read_reg mstatus_ref) \ ((\ (w__0 :: Mstatus) . read_reg misa_ref \ ((\ (w__1 :: Misa) . (write_reg mstatus_ref ((set_mstatus_SXL w__0 ((get_Misa_MXL w__1 :: 2 Word.word)))) \ read_reg mstatus_ref) \ ((\ (w__2 :: Mstatus) . read_reg misa_ref \ ((\ (w__3 :: Misa) . (((((((((((((((((write_reg mstatus_ref ((set_mstatus_UXL w__2 ((get_Misa_MXL w__3 :: 2 Word.word)))) \ set_Mstatus_SD mstatus_ref ( 0b0 :: 1 Word.word)) \ set_Mstatush_bits mstatush_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ set_Minterrupts_bits mip_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ set_Minterrupts_bits mie_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ set_Minterrupts_bits mideleg_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ set_Medeleg_bits medeleg_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ set_Mtvec_bits mtvec_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ set_Mcause_bits mcause_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ write_reg mepc_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ write_reg mtval_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ write_reg mscratch_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ write_reg mcycle_ref ((EXTZ (( 64 :: int)::ii) ( 0b0 :: 1 Word.word) :: 64 Word.word))) \ write_reg mtime_ref ((EXTZ (( 64 :: int)::ii) ( 0b0 :: 1 Word.word) :: 64 Word.word))) \ set_Counteren_bits mcounteren_ref ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word))) \ write_reg minstret_ref ((EXTZ (( 64 :: int)::ii) ( 0b0 :: 1 Word.word) :: 64 Word.word))) \ write_reg minstret_written_ref False) \ init_pmp () ) \ (if ((get_config_print_reg () )) then read_reg mstatus_ref \ ((\ (w__4 :: Mstatus) . return ((print_dbg (((@) (''CSR mstatus <- '') (((@) ((string_of_bits ((get_Mstatus_bits w__4 :: 32 Word.word)))) (((@) ('' (input: '') (((@) ((string_of_bits ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word)))) ('')''))))))))))))) else return () ))))))))))\ \ \\val MemoryOpResult_add_meta : forall 't. MemoryOpResult 't -> unit -> MemoryOpResult (('t * unit))\\ fun MemoryOpResult_add_meta :: \ 't MemoryOpResult \ unit \('t*unit)MemoryOpResult \ where \ MemoryOpResult_add_meta (MemValue (v)) m = ( MemValue (v, m))\ for v :: " 't " and m :: " unit " |\ MemoryOpResult_add_meta (MemException (e)) m = ( MemException e )\ for e :: " ExceptionType " and m :: " unit " \ \\val MemoryOpResult_drop_meta : forall 't. MemoryOpResult (('t * unit)) -> MemoryOpResult 't\\ fun MemoryOpResult_drop_meta :: \('t*unit)MemoryOpResult \ 't MemoryOpResult \ where \ MemoryOpResult_drop_meta (MemValue ((v, m))) = ( MemValue v )\ for m :: " unit " and v :: " 't " |\ MemoryOpResult_drop_meta (MemException (e)) = ( MemException e )\ for e :: " ExceptionType " \ \\val elf_tohost : unit -> ii\\ \ \\val elf_entry : unit -> ii\\ \ \\val phys_mem_segments : unit -> list ((mword ty32 * mword ty32))\\ definition phys_mem_segments :: \ unit \((32)Word.word*(32)Word.word)list \ where \ phys_mem_segments _ = ( ((plat_rom_base () :: 32 Word.word), (plat_rom_size () :: 32 Word.word)) # (((plat_ram_base () :: 32 Word.word), (plat_ram_size () :: 32 Word.word)) # []))\ \ \\val within_phys_mem : mword ty32 -> integer -> bool\\ definition within_phys_mem :: \(32)Word.word \ int \ bool \ where \ within_phys_mem (addr :: xlenbits) (width :: int) = ( (let addr_int = (Word.uint addr) in (let ram_base_int = (Word.uint ((plat_ram_base () :: 32 Word.word))) in (let rom_base_int = (Word.uint ((plat_rom_base () :: 32 Word.word))) in (let ram_size_int = (Word.uint ((plat_ram_size () :: 32 Word.word))) in (let rom_size_int = (Word.uint ((plat_rom_size () :: 32 Word.word))) in if (((((ram_base_int \ addr_int)) \ ((((addr_int + ((id0 width)))) \ ((ram_base_int + ram_size_int))))))) then True else if (((((rom_base_int \ addr_int)) \ ((((addr_int + ((id0 width)))) \ ((rom_base_int + rom_size_int))))))) then True else (let (_ :: unit) = (print_dbg (((@) (''within_phys_mem: '') (((@) ((string_of_bits addr)) ('' not within phys-mem:'')))))) in (let (_ :: unit) = (print_dbg (((@) ('' plat_rom_base: '') ((string_of_bits ((plat_rom_base () :: 32 Word.word))))))) in (let (_ :: unit) = (print_dbg (((@) ('' plat_rom_size: '') ((string_of_bits ((plat_rom_size () :: 32 Word.word))))))) in (let (_ :: unit) = (print_dbg (((@) ('' plat_ram_base: '') ((string_of_bits ((plat_ram_base () :: 32 Word.word))))))) in (let (_ :: unit) = (print_dbg (((@) ('' plat_ram_size: '') ((string_of_bits ((plat_ram_size () :: 32 Word.word))))))) in False)))))))))))\ for addr :: "(32)Word.word " and width :: " int " \ \\val within_clint : mword ty32 -> integer -> bool\\ definition within_clint :: \(32)Word.word \ int \ bool \ where \ within_clint (addr :: xlenbits) (width :: int) = ( (let addr_int = (Word.uint addr) in (let clint_base_int = (Word.uint ((plat_clint_base () :: 32 Word.word))) in (let clint_size_int = (Word.uint ((plat_clint_size () :: 32 Word.word))) in (((clint_base_int \ addr_int)) \ ((((addr_int + ((id0 width)))) \ ((clint_base_int + clint_size_int)))))))))\ for addr :: "(32)Word.word " and width :: " int " \ \\val within_htif_writable : mword ty32 -> integer -> bool\\ definition within_htif_writable :: \(32)Word.word \ int \ bool \ where \ within_htif_writable (addr :: xlenbits) (width :: int) = ( ((((((plat_htif_tohost () :: 32 Word.word)) = addr))) \ ((((((((add_vec_int ((plat_htif_tohost () :: 32 Word.word)) (( 4 :: int)::ii) :: 32 Word.word)) = addr))) \ (((width = (( 4 :: int)::ii)))))))))\ for addr :: "(32)Word.word " and width :: " int " \ \\val within_htif_readable : mword ty32 -> integer -> bool\\ definition within_htif_readable :: \(32)Word.word \ int \ bool \ where \ within_htif_readable (addr :: xlenbits) (width :: int) = ( ((((((plat_htif_tohost () :: 32 Word.word)) = addr))) \ ((((((((add_vec_int ((plat_htif_tohost () :: 32 Word.word)) (( 4 :: int)::ii) :: 32 Word.word)) = addr))) \ (((width = (( 4 :: int)::ii)))))))))\ for addr :: "(32)Word.word " and width :: " int " definition MSIP_BASE :: \(32)Word.word \ where \ MSIP_BASE = ( ( 0x00000000 :: 32 Word.word))\ definition MTIMECMP_BASE :: \(32)Word.word \ where \ MTIMECMP_BASE = ( ( 0x00004000 :: 32 Word.word))\ definition MTIMECMP_BASE_HI :: \(32)Word.word \ where \ MTIMECMP_BASE_HI = ( ( 0x00004004 :: 32 Word.word))\ definition MTIME_BASE :: \(32)Word.word \ where \ MTIME_BASE = ( ( 0x0000BFF8 :: 32 Word.word))\ definition MTIME_BASE_HI :: \(32)Word.word \ where \ MTIME_BASE_HI = ( ( 0x0000BFFC :: 32 Word.word))\ \ \\val clint_load : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> mword ty32 -> integer -> M (MemoryOpResult (mword 'int8_times_n))\\ definition clint_load :: \(unit)AccessType \(32)Word.word \ int \((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad \ where \ clint_load t addr width = ( (let addr = ((sub_vec addr ((plat_clint_base () :: 32 Word.word)) :: 32 Word.word)) in if ((((((addr = MSIP_BASE))) \ ((((((((id0 width)) = (( 8 :: int)::ii)))) \ (((((id0 width)) = (( 4 :: int)::ii)))))))))) then ((if ((get_config_print_platform () )) then read_reg mip_ref \ ((\ (w__0 :: Minterrupts) . return ((print_dbg (((@) (''clint['') (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits ((get_Minterrupts_MSI w__0 :: 1 Word.word)))))))))))))) else return () ) \ read_reg mip_ref) \ ((\ (w__1 :: Minterrupts) . return (MemValue ((zero_extend ((get_Minterrupts_MSI w__1 :: 1 Word.word)) (((( 8 :: int)::ii) * ((id0 width)))) :: ( 'int8_times_n::len)Word.word))))) else if ((((((addr = MTIMECMP_BASE))) \ (((((id0 width)) = (( 4 :: int)::ii))))))) then ((if ((get_config_print_platform () )) then (read_reg mtimecmp_ref :: ( 64 Word.word) M) \ ((\ (w__2 :: 64 Word.word) . return ((print_dbg (((@) (''clint<4>['') (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits ((subrange_vec_dec w__2 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)))))))))))))) else return () ) \ (read_reg mtimecmp_ref :: ( 64 Word.word) M)) \ ((\ (w__3 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend ((subrange_vec_dec w__3 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) (( 32 :: int)::ii) :: 32 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else if ((((((addr = MTIMECMP_BASE))) \ (((((id0 width)) = (( 8 :: int)::ii))))))) then ((if ((get_config_print_platform () )) then (read_reg mtimecmp_ref :: ( 64 Word.word) M) \ ((\ (w__4 :: 64 Word.word) . return ((print_dbg (((@) (''clint<8>['') (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits w__4)))))))))))) else return () ) \ (read_reg mtimecmp_ref :: ( 64 Word.word) M)) \ ((\ (w__5 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend w__5 (( 64 :: int)::ii) :: 64 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else if ((((((addr = MTIMECMP_BASE_HI))) \ (((((id0 width)) = (( 4 :: int)::ii))))))) then ((if ((get_config_print_platform () )) then (read_reg mtimecmp_ref :: ( 64 Word.word) M) \ ((\ (w__6 :: 64 Word.word) . return ((print_dbg (((@) (''clint-hi<4>['') (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits ((subrange_vec_dec w__6 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)))))))))))))) else return () ) \ (read_reg mtimecmp_ref :: ( 64 Word.word) M)) \ ((\ (w__7 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend ((subrange_vec_dec w__7 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) (( 32 :: int)::ii) :: 32 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else if ((((((addr = MTIME_BASE))) \ (((((id0 width)) = (( 4 :: int)::ii))))))) then ((if ((get_config_print_platform () )) then (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__8 :: 64 Word.word) . return ((print_dbg (((@) (''clint['') (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits w__8)))))))))))) else return () ) \ (read_reg mtime_ref :: ( 64 Word.word) M)) \ ((\ (w__9 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend ((subrange_vec_dec w__9 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) (( 32 :: int)::ii) :: 32 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else if ((((((addr = MTIME_BASE))) \ (((((id0 width)) = (( 8 :: int)::ii))))))) then ((if ((get_config_print_platform () )) then (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__10 :: 64 Word.word) . return ((print_dbg (((@) (''clint['') (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits w__10)))))))))))) else return () ) \ (read_reg mtime_ref :: ( 64 Word.word) M)) \ ((\ (w__11 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend w__11 (( 64 :: int)::ii) :: 64 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else if ((((((addr = MTIME_BASE_HI))) \ (((((id0 width)) = (( 4 :: int)::ii))))))) then ((if ((get_config_print_platform () )) then (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__12 :: 64 Word.word) . return ((print_dbg (((@) (''clint['') (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits w__12)))))))))))) else return () ) \ (read_reg mtime_ref :: ( 64 Word.word) M)) \ ((\ (w__13 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend ((subrange_vec_dec w__13 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) (( 32 :: int)::ii) :: 32 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''clint['') (((@) ((string_of_bits addr)) (''] -> ''))))) else () ) in return ((case t of Execute (_) => MemException (E_Fetch_Access_Fault () ) | Read (Data1) => MemException (E_Load_Access_Fault () ) | _ => MemException (E_SAMO_Access_Fault () ) )))))\ for t :: "(unit)AccessType " and addr :: "(32)Word.word " and width :: " int " \ \\val clint_dispatch : unit -> M unit\\ definition clint_dispatch :: \ unit \((register_value),(unit),(exception))monad \ where \ clint_dispatch _ = ( (((if ((get_config_print_platform () )) then (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . return ((print_dbg (((@) (''clint::tick mtime <- '') ((string_of_bits w__0)))))))) else return () ) \ set_Minterrupts_MTI mip_ref ( 0b0 :: 1 Word.word)) \ (read_reg mtimecmp_ref :: ( 64 Word.word) M)) \ ((\ (w__1 :: 64 Word.word) . (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__2 :: 64 Word.word) . if ((zopz0zIzJ_u w__1 w__2)) then (if ((get_config_print_platform () )) then (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__3 :: 64 Word.word) . return ((print_dbg (((@) ('' clint timer pending at mtime '') ((string_of_bits w__3)))))))) else return () ) \ set_Minterrupts_MTI mip_ref ( 0b1 :: 1 Word.word) else return () )))))\ \ \\val clint_store : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)\\ definition clint_store :: \(32)Word.word \ int \('int8_times_n::len)Word.word \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ clint_store addr width data = ( (let addr = ((sub_vec addr ((plat_clint_base () :: 32 Word.word)) :: 32 Word.word)) in if ((((((addr = MSIP_BASE))) \ ((((((((id0 width)) = (( 8 :: int)::ii)))) \ (((((id0 width)) = (( 4 :: int)::ii)))))))))) then (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''clint['') (((@) ((string_of_bits addr)) (((@) (''] <- '') (((@) ((string_of_bits data)) (((@) ('' (mip.MSI <- '') (((@) ((string_of_bit ((access_vec_dec data (( 0 :: int)::ii))))) ('')''))))))))))))) else () ) in (set_Minterrupts_MSI mip_ref (vec_of_bits [access_vec_dec data (( 0 :: int)::ii)] :: 1 Word.word) \ clint_dispatch () ) \ return (MemValue True)) else if ((((((addr = MTIMECMP_BASE))) \ (((((id0 width)) = (( 8 :: int)::ii))))))) then (let (data :: 64 Word.word) = ((Word.ucast data :: 64 Word.word)) in (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''clint<8>['') (((@) ((string_of_bits addr)) (((@) (''] <- '') (((@) ((string_of_bits data)) ('' (mtimecmp)''))))))))) else () ) in (write_reg mtimecmp_ref ((zero_extend data (( 64 :: int)::ii) :: 64 Word.word)) \ clint_dispatch () ) \ return (MemValue True))) else if ((((((addr = MTIMECMP_BASE))) \ (((((id0 width)) = (( 4 :: int)::ii))))))) then (let (data :: 32 Word.word) = ((Word.ucast data :: 32 Word.word)) in (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''clint<4>['') (((@) ((string_of_bits addr)) (((@) (''] <- '') (((@) ((string_of_bits data)) ('' (mtimecmp)''))))))))) else () ) in (read_reg mtimecmp_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (write_reg mtimecmp_ref ((update_subrange_vec_dec w__0 (( 31 :: int)::ii) (( 0 :: int)::ii) ((zero_extend data (( 32 :: int)::ii) :: 32 Word.word)) :: 64 Word.word)) \ clint_dispatch () ) \ return (MemValue True))))) else if ((((((addr = MTIMECMP_BASE_HI))) \ (((((id0 width)) = (( 4 :: int)::ii))))))) then (let (data :: 32 Word.word) = ((Word.ucast data :: 32 Word.word)) in (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''clint<4>['') (((@) ((string_of_bits addr)) (((@) (''] <- '') (((@) ((string_of_bits data)) ('' (mtimecmp)''))))))))) else () ) in (read_reg mtimecmp_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . (write_reg mtimecmp_ref ((update_subrange_vec_dec w__1 (( 63 :: int)::ii) (( 32 :: int)::ii) ((zero_extend data (( 32 :: int)::ii) :: 32 Word.word)) :: 64 Word.word)) \ clint_dispatch () ) \ return (MemValue True))))) else (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''clint['') (((@) ((string_of_bits addr)) (((@) (''] <- '') (((@) ((string_of_bits data)) ('' ()''))))))))) else () ) in return (MemException (E_SAMO_Access_Fault () )))))\ for addr :: "(32)Word.word " and width :: " int " and data :: "('int8_times_n::len)Word.word " \ \\val tick_clock : unit -> M unit\\ definition tick_clock :: \ unit \((register_value),(unit),(exception))monad \ where \ tick_clock _ = ( read_reg mcountinhibit_ref \ ((\ (w__0 :: Counterin) . ((if (((((get_Counterin_CY w__0 :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then (read_reg mcycle_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . write_reg mcycle_ref ((add_vec_int w__1 (( 1 :: int)::ii) :: 64 Word.word)))) else return () ) \ (read_reg mtime_ref :: ( 64 Word.word) M)) \ ((\ (w__2 :: 64 Word.word) . write_reg mtime_ref ((add_vec_int w__2 (( 1 :: int)::ii) :: 64 Word.word)) \ clint_dispatch () )))))\ \ \\val undefined_htif_cmd : unit -> M htif_cmd\\ definition undefined_htif_cmd :: \ unit \((register_value),(htif_cmd),(exception))monad \ where \ undefined_htif_cmd _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . return ((| htif_cmd_bits = w__0 |)))))\ \ \\val Mk_htif_cmd : mword ty64 -> htif_cmd\\ definition Mk_htif_cmd :: \(64)Word.word \ htif_cmd \ where \ Mk_htif_cmd v = ( (| htif_cmd_bits = v |) )\ for v :: "(64)Word.word " definition get_htif_cmd_bits :: \ htif_cmd \(64)Word.word \ where \ get_htif_cmd_bits v = ( (subrange_vec_dec(htif_cmd_bits v) (( 63 :: int)::ii) (( 0 :: int)::ii) :: 64 Word.word))\ for v :: " htif_cmd " definition set_htif_cmd_bits :: \((regstate),(register_value),(htif_cmd))register_ref \(64)Word.word \((register_value),(unit),(exception))monad \ where \ set_htif_cmd_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| htif_cmd_bits := ((update_subrange_vec_dec(htif_cmd_bits r) (( 63 :: int)::ii) (( 0 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(htif_cmd))register_ref " and v :: "(64)Word.word " definition update_htif_cmd_bits :: \ htif_cmd \(64)Word.word \ htif_cmd \ where \ update_htif_cmd_bits v x = ( ( v (| htif_cmd_bits := ((update_subrange_vec_dec(htif_cmd_bits v) (( 63 :: int)::ii) (( 0 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " htif_cmd " and x :: "(64)Word.word " \ \\val _get_htif_cmd_cmd : htif_cmd -> mword ty8\\ definition get_htif_cmd_cmd :: \ htif_cmd \(8)Word.word \ where \ get_htif_cmd_cmd v = ( (subrange_vec_dec(htif_cmd_bits v) (( 55 :: int)::ii) (( 48 :: int)::ii) :: 8 Word.word))\ for v :: " htif_cmd " \ \\val _set_htif_cmd_cmd : register_ref regstate register_value htif_cmd -> mword ty8 -> M unit\\ definition set_htif_cmd_cmd :: \((regstate),(register_value),(htif_cmd))register_ref \(8)Word.word \((register_value),(unit),(exception))monad \ where \ set_htif_cmd_cmd r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| htif_cmd_bits := ((update_subrange_vec_dec(htif_cmd_bits r) (( 55 :: int)::ii) (( 48 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(htif_cmd))register_ref " and v :: "(8)Word.word " \ \\val _update_htif_cmd_cmd : htif_cmd -> mword ty8 -> htif_cmd\\ definition update_htif_cmd_cmd :: \ htif_cmd \(8)Word.word \ htif_cmd \ where \ update_htif_cmd_cmd v x = ( ( v (| htif_cmd_bits := ((update_subrange_vec_dec(htif_cmd_bits v) (( 55 :: int)::ii) (( 48 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " htif_cmd " and x :: "(8)Word.word " \ \\val _get_htif_cmd_device : htif_cmd -> mword ty8\\ definition get_htif_cmd_device :: \ htif_cmd \(8)Word.word \ where \ get_htif_cmd_device v = ( (subrange_vec_dec(htif_cmd_bits v) (( 63 :: int)::ii) (( 56 :: int)::ii) :: 8 Word.word))\ for v :: " htif_cmd " \ \\val _set_htif_cmd_device : register_ref regstate register_value htif_cmd -> mword ty8 -> M unit\\ definition set_htif_cmd_device :: \((regstate),(register_value),(htif_cmd))register_ref \(8)Word.word \((register_value),(unit),(exception))monad \ where \ set_htif_cmd_device r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| htif_cmd_bits := ((update_subrange_vec_dec(htif_cmd_bits r) (( 63 :: int)::ii) (( 56 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(htif_cmd))register_ref " and v :: "(8)Word.word " \ \\val _update_htif_cmd_device : htif_cmd -> mword ty8 -> htif_cmd\\ definition update_htif_cmd_device :: \ htif_cmd \(8)Word.word \ htif_cmd \ where \ update_htif_cmd_device v x = ( ( v (| htif_cmd_bits := ((update_subrange_vec_dec(htif_cmd_bits v) (( 63 :: int)::ii) (( 56 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " htif_cmd " and x :: "(8)Word.word " \ \\val _get_htif_cmd_payload : htif_cmd -> mword ty48\\ definition get_htif_cmd_payload :: \ htif_cmd \(48)Word.word \ where \ get_htif_cmd_payload v = ( (subrange_vec_dec(htif_cmd_bits v) (( 47 :: int)::ii) (( 0 :: int)::ii) :: 48 Word.word))\ for v :: " htif_cmd " \ \\val _set_htif_cmd_payload : register_ref regstate register_value htif_cmd -> mword ty48 -> M unit\\ definition set_htif_cmd_payload :: \((regstate),(register_value),(htif_cmd))register_ref \(48)Word.word \((register_value),(unit),(exception))monad \ where \ set_htif_cmd_payload r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| htif_cmd_bits := ((update_subrange_vec_dec(htif_cmd_bits r) (( 47 :: int)::ii) (( 0 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(htif_cmd))register_ref " and v :: "(48)Word.word " \ \\val _update_htif_cmd_payload : htif_cmd -> mword ty48 -> htif_cmd\\ definition update_htif_cmd_payload :: \ htif_cmd \(48)Word.word \ htif_cmd \ where \ update_htif_cmd_payload v x = ( ( v (| htif_cmd_bits := ((update_subrange_vec_dec(htif_cmd_bits v) (( 47 :: int)::ii) (( 0 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " htif_cmd " and x :: "(48)Word.word " \ \\val reset_htif : unit -> M unit\\ definition reset_htif :: \ unit \((register_value),(unit),(exception))monad \ where \ reset_htif _ = ( (write_reg htif_cmd_write_ref B0 \ write_reg htif_payload_writes_ref ( 0x0 :: 4 Word.word)) \ write_reg htif_tohost_ref ((EXTZ (( 64 :: int)::ii) ( 0b0 :: 1 Word.word) :: 64 Word.word)))\ \ \\val htif_load : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> mword ty32 -> integer -> M (MemoryOpResult (mword 'int8_times_n))\\ definition htif_load :: \(unit)AccessType \(32)Word.word \ int \((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad \ where \ htif_load t paddr width = ( (if ((get_config_print_platform () )) then (read_reg htif_tohost_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . return ((print_dbg (((@) (''htif['') (((@) ((string_of_bits paddr)) (((@) (''] -> '') ((string_of_bits w__0)))))))))))) else return () ) \ (if ((((((width = (( 8 :: int)::ii)))) \ (((paddr = ((plat_htif_tohost () :: 32 Word.word)))))))) then (read_reg htif_tohost_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend w__1 (( 64 :: int)::ii) :: 64 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else if ((((((width = (( 4 :: int)::ii)))) \ (((paddr = ((plat_htif_tohost () :: 32 Word.word)))))))) then (read_reg htif_tohost_ref :: ( 64 Word.word) M) \ ((\ (w__2 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend ((subrange_vec_dec w__2 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) (( 32 :: int)::ii) :: 32 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else if ((((((width = (( 4 :: int)::ii)))) \ (((paddr = ((add_vec_int ((plat_htif_tohost () :: 32 Word.word)) (( 4 :: int)::ii) :: 32 Word.word)))))))) then (read_reg htif_tohost_ref :: ( 64 Word.word) M) \ ((\ (w__3 :: 64 Word.word) . return (MemValue ((Word.ucast ((zero_extend ((subrange_vec_dec w__3 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) (( 32 :: int)::ii) :: 32 Word.word)) :: ( 'int8_times_n::len)Word.word))))) else return ((case t of Execute (_) => MemException (E_Fetch_Access_Fault () ) | Read (Data1) => MemException (E_Load_Access_Fault () ) | _ => MemException (E_SAMO_Access_Fault () ) ))))\ for t :: "(unit)AccessType " and paddr :: "(32)Word.word " and width :: " int " \ \\val htif_store : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)\\ definition htif_store :: \(32)Word.word \ int \('int8_times_n::len)Word.word \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ htif_store paddr width data = ( (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''htif['') (((@) ((string_of_bits paddr)) (((@) (''] <- '') ((string_of_bits data)))))))) else () ) in ((if (((width = (( 8 :: int)::ii)))) then (let (data :: 64 Word.word) = ((Word.ucast data :: 64 Word.word)) in (write_reg htif_cmd_write_ref B1 \ (read_reg htif_payload_writes_ref :: ( 4 Word.word) M)) \ ((\ (w__0 :: 4 Word.word) . write_reg htif_payload_writes_ref ((add_vec_int w__0 (( 1 :: int)::ii) :: 4 Word.word)) \ write_reg htif_tohost_ref ((EXTZ (( 64 :: int)::ii) data :: 64 Word.word))))) else if ((((((width = (( 4 :: int)::ii)))) \ (((paddr = ((plat_htif_tohost () :: 32 Word.word)))))))) then (let (data :: 32 Word.word) = ((Word.ucast data :: 32 Word.word)) in (read_reg htif_tohost_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . ((if (((data = ((subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))))) then (read_reg htif_payload_writes_ref :: ( 4 Word.word) M) \ ((\ (w__2 :: 4 Word.word) . write_reg htif_payload_writes_ref ((add_vec_int w__2 (( 1 :: int)::ii) :: 4 Word.word)))) else write_reg htif_payload_writes_ref ( 0x1 :: 4 Word.word)) \ (read_reg htif_tohost_ref :: ( 64 Word.word) M)) \ ((\ (w__3 :: 64 Word.word) . write_reg htif_tohost_ref ((update_subrange_vec_dec w__3 (( 31 :: int)::ii) (( 0 :: int)::ii) data :: 64 Word.word))))))) else if ((((((width = (( 4 :: int)::ii)))) \ (((paddr = ((add_vec_int ((plat_htif_tohost () :: 32 Word.word)) (( 4 :: int)::ii) :: 32 Word.word)))))))) then (let (data :: 32 Word.word) = ((Word.ucast data :: 32 Word.word)) in (read_reg htif_tohost_ref :: ( 64 Word.word) M) \ ((\ (w__4 :: 64 Word.word) . (((if (((((subrange_vec_dec data (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = ((subrange_vec_dec w__4 (( 47 :: int)::ii) (( 32 :: int)::ii) :: 16 Word.word))))) then (read_reg htif_payload_writes_ref :: ( 4 Word.word) M) \ ((\ (w__5 :: 4 Word.word) . write_reg htif_payload_writes_ref ((add_vec_int w__5 (( 1 :: int)::ii) :: 4 Word.word)))) else write_reg htif_payload_writes_ref ( 0x1 :: 4 Word.word)) \ write_reg htif_cmd_write_ref B1) \ (read_reg htif_tohost_ref :: ( 64 Word.word) M)) \ ((\ (w__6 :: 64 Word.word) . write_reg htif_tohost_ref ((update_subrange_vec_dec w__6 (( 63 :: int)::ii) (( 32 :: int)::ii) data :: 64 Word.word))))))) else write_reg htif_tohost_ref ((EXTZ (( 64 :: int)::ii) data :: 64 Word.word))) \ or_boolM (and_boolM (read_reg htif_cmd_write_ref \ ((\ (w__7 :: bitU) . return (((w__7 = B1)))))) ((read_reg htif_payload_writes_ref :: ( 4 Word.word) M) \ ((\ (w__8 :: 4 Word.word) . return ((((Word.uint w__8)) > (( 0 :: int)::ii))))))) ((read_reg htif_payload_writes_ref :: ( 4 Word.word) M) \ ((\ (w__10 :: 4 Word.word) . return ((((Word.uint w__10)) > (( 2 :: int)::ii))))))) \ ((\ (w__11 :: bool) . (if w__11 then (read_reg htif_tohost_ref :: ( 64 Word.word) M) \ ((\ (w__12 :: 64 Word.word) . (let cmd = (Mk_htif_cmd w__12) in (let b__0 = ((get_htif_cmd_device cmd :: 8 Word.word)) in if (((b__0 = ( 0x00 :: 8 Word.word)))) then (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''htif-syscall-proxy cmd: '') ((string_of_bits ((get_htif_cmd_payload cmd :: 48 Word.word)))))) else () ) in if (((((access_vec_dec ((get_htif_cmd_payload cmd :: 48 Word.word)) (( 0 :: int)::ii))) = B1))) then write_reg htif_done_ref True \ write_reg htif_exit_code_ref ((shiftr ((zero_extend ((get_htif_cmd_payload cmd :: 48 Word.word)) (( 64 :: int)::ii) :: 64 Word.word)) (( 1 :: int)::ii) :: 64 Word.word)) else return () ) else if (((b__0 = ( 0x01 :: 8 Word.word)))) then (let (_ :: unit) = (if ((get_config_print_platform () )) then print_dbg (((@) (''htif-term cmd: '') ((string_of_bits ((get_htif_cmd_payload cmd :: 48 Word.word)))))) else () ) in (let b__2 = ((get_htif_cmd_cmd cmd :: 8 Word.word)) in (let (_ :: unit) = (if (((b__2 = ( 0x00 :: 8 Word.word)))) then () else if (((b__2 = ( 0x01 :: 8 Word.word)))) then plat_term_write ((subrange_vec_dec ((get_htif_cmd_payload cmd :: 48 Word.word)) (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)) else print_endline (((@) (''Unknown term cmd: '') ((string_of_bits b__2))))) in reset_htif () ))) else return ((print_endline (((@) (''htif-???? cmd: '') ((string_of_bits data)))))))))) else return () ) \ return (MemValue True)))))\ for paddr :: "(32)Word.word " and width :: " int " and data :: "('int8_times_n::len)Word.word " \ \\val htif_tick : unit -> M unit\\ definition htif_tick :: \ unit \((register_value),(unit),(exception))monad \ where \ htif_tick _ = ( ((if ((get_config_print_platform () )) then (read_reg htif_tohost_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . return ((print_dbg (((@) (''htif::tick '') ((string_of_bits w__0)))))))) else return () ) \ (read_reg htif_tohost_ref :: ( 64 Word.word) M)) \ ((\ (w__1 :: 64 Word.word) . write_reg htif_tohost_ref w__1)))\ \ \\val within_mmio_readable : mword ty32 -> integer -> bool\\ definition within_mmio_readable :: \(32)Word.word \ int \ bool \ where \ within_mmio_readable (addr :: xlenbits) (width :: int) = ( (((within_clint addr width)) \ (((((within_htif_readable addr width)) \ (((( 1 :: int)::ii) \ ((id0 width)))))))))\ for addr :: "(32)Word.word " and width :: " int " \ \\val within_mmio_writable : mword ty32 -> integer -> bool\\ definition within_mmio_writable :: \(32)Word.word \ int \ bool \ where \ within_mmio_writable (addr :: xlenbits) (width :: int) = ( (((within_clint addr width)) \ (((((within_htif_writable addr width)) \ ((((id0 width)) \ (( 8 :: int)::ii))))))))\ for addr :: "(32)Word.word " and width :: " int " \ \\val mmio_read : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> mword ty32 -> integer -> M (MemoryOpResult (mword 'int8_times_n))\\ definition mmio_read :: \(ext_access_type)AccessType \(32)Word.word \ int \((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad \ where \ mmio_read (t :: ext_access_type AccessType) (paddr :: xlenbits) (width :: int) = ( if ((within_clint paddr width)) then (clint_load t paddr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) else if (((((within_htif_readable paddr width)) \ (((( 1 :: int)::ii) \ ((id0 width))))))) then (htif_load t paddr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) else return ((case t of Execute (_) => MemException (E_Fetch_Access_Fault () ) | Read (Data1) => MemException (E_Load_Access_Fault () ) | _ => MemException (E_SAMO_Access_Fault () ) )))\ for t :: "(ext_access_type)AccessType " and paddr :: "(32)Word.word " and width :: " int " \ \\val mmio_write : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)\\ definition mmio_write :: \(32)Word.word \ int \('int8_times_n::len)Word.word \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ mmio_write (paddr :: xlenbits) (width :: int) (data :: ( 'int8_times_n::len)Word.word) = ( if ((within_clint paddr width)) then clint_store paddr width data else if (((((within_htif_writable paddr width)) \ ((((id0 width)) \ (( 8 :: int)::ii)))))) then htif_store paddr width data else return (MemException (E_SAMO_Access_Fault () )))\ for paddr :: "(32)Word.word " and width :: " int " and data :: "('int8_times_n::len)Word.word " \ \\val init_platform : unit -> M unit\\ definition init_platform :: \ unit \((register_value),(unit),(exception))monad \ where \ init_platform _ = ( (((write_reg htif_tohost_ref ((EXTZ (( 64 :: int)::ii) ( 0b0 :: 1 Word.word) :: 64 Word.word)) \ write_reg htif_done_ref False) \ write_reg htif_exit_code_ref ((EXTZ (( 64 :: int)::ii) ( 0b0 :: 1 Word.word) :: 64 Word.word))) \ write_reg htif_cmd_write_ref B0) \ write_reg htif_payload_writes_ref ((EXTZ (( 4 :: int)::ii) ( 0b0 :: 1 Word.word) :: 4 Word.word)))\ \ \\val tick_platform : unit -> M unit\\ definition tick_platform :: \ unit \((register_value),(unit),(exception))monad \ where \ tick_platform _ = ( htif_tick () )\ \ \\val handle_illegal : unit -> M unit\\ definition handle_illegal :: \ unit \((register_value),(unit),(exception))monad \ where \ handle_illegal _ = ( (if ((plat_mtval_has_illegal_inst_bits () )) then (read_reg instbits_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return (Some w__0))) else return None) \ ((\ info . (let (t :: sync_exception) = ((| sync_exception_trap = (E_Illegal_Instr () ), sync_exception_excinfo = info, sync_exception_ext_exception = None |)) in read_reg cur_privilege_ref \ ((\ (w__1 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . (exception_handler w__1 (CTL_TRAP t) w__2 :: ( 32 Word.word) M) \ ((\ (w__3 :: 32 Word.word) . set_next_pc w__3))))))))))\ \ \\val platform_wfi : unit -> M unit\\ definition platform_wfi :: \ unit \((register_value),(unit),(exception))monad \ where \ platform_wfi _ = ( (let (_ :: unit) = (cancel_reservation () ) in (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . (read_reg mtimecmp_ref :: ( 64 Word.word) M) \ ((\ (w__1 :: 64 Word.word) . if ((zopz0zI_u w__0 w__1)) then (read_reg mtimecmp_ref :: ( 64 Word.word) M) \ ((\ (w__2 :: 64 Word.word) . (write_reg mtime_ref w__2 \ (read_reg mtimecmp_ref :: ( 64 Word.word) M)) \ ((\ (w__3 :: 64 Word.word) . write_reg mcycle_ref w__3)))) else return () ))))))\ \ \\val is_aligned_addr : mword ty32 -> integer -> bool\\ definition is_aligned_addr :: \(32)Word.word \ int \ bool \ where \ is_aligned_addr (addr :: xlenbits) (width :: int) = ( (((((Word.uint addr)) mod width)) = (( 0 :: int)::ii)))\ for addr :: "(32)Word.word " and width :: " int " \ \\val read_kind_of_flags : bool -> bool -> bool -> maybe read_kind\\ fun read_kind_of_flags :: \ bool \ bool \ bool \(read_kind)option \ where \ read_kind_of_flags (False :: bool) (False :: bool) (False :: bool) = ( Some Read_plain )\ |\ read_kind_of_flags (True :: bool) (False :: bool) (False :: bool) = ( Some Read_RISCV_acquire )\ |\ read_kind_of_flags (True :: bool) (True :: bool) (False :: bool) = ( Some Read_RISCV_strong_acquire )\ |\ read_kind_of_flags (False :: bool) (False :: bool) (True :: bool) = ( Some Read_RISCV_reserved )\ |\ read_kind_of_flags (True :: bool) (False :: bool) (True :: bool) = ( Some Read_RISCV_reserved_acquire )\ |\ read_kind_of_flags (True :: bool) (True :: bool) (True :: bool) = ( Some Read_RISCV_reserved_strong_acquire )\ |\ read_kind_of_flags (False :: bool) (True :: bool) (False :: bool) = ( None )\ |\ read_kind_of_flags (False :: bool) (True :: bool) (True :: bool) = ( None )\ \ \\val phys_mem_read : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> mword ty32 -> integer -> bool -> bool -> bool -> bool -> M (MemoryOpResult ((mword 'int8_times_n * unit)))\\ definition phys_mem_read :: \(ext_access_type)AccessType \(32)Word.word \ int \ bool \ bool \ bool \ bool \((register_value),((('int8_times_n::len)Word.word*unit)MemoryOpResult),(exception))monad \ where \ phys_mem_read (t :: ext_access_type AccessType) (paddr :: xlenbits) (width :: int) (aq :: bool) (rl :: bool) (res :: bool) (meta :: bool) = ( (case ((read_kind_of_flags aq rl res)) of Some (rk) => (read_ram rk paddr width meta :: ((( 'int8_times_n::len)Word.word * unit)) M) \ ((\ (w__0 :: (( 'int8_times_n::len)Word.word * unit)) . return (Some w__0))) | None => return None ) \ ((\ (w__1 :: ((( 'int8_times_n::len)Word.word * unit))option) . (let result = w__1 in return ((case (t, result) of (Execute (_), None) => MemException (E_Fetch_Access_Fault () ) | (Read (Data1), None) => MemException (E_Load_Access_Fault () ) | (_, None) => MemException (E_SAMO_Access_Fault () ) | (_, Some ((v, m))) => (let (_ :: unit) = (if ((get_config_print_mem () )) then print_dbg (((@) (''mem['') (((@) ((accessType_to_str t)) (((@) ('','') (((@) ((string_of_bits paddr)) (((@) (''] -> '') ((string_of_bits v)))))))))))) else () ) in MemValue (v, m)) ))))))\ for t :: "(ext_access_type)AccessType " and paddr :: "(32)Word.word " and width :: " int " and aq :: " bool " and rl :: " bool " and res :: " bool " and meta :: " bool " \ \\val checked_mem_read : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> mword ty32 -> integer -> bool -> bool -> bool -> bool -> M (MemoryOpResult ((mword 'int8_times_n * unit)))\\ definition checked_mem_read :: \(ext_access_type)AccessType \(32)Word.word \ int \ bool \ bool \ bool \ bool \((register_value),((('int8_times_n::len)Word.word*unit)MemoryOpResult),(exception))monad \ where \ checked_mem_read (t :: ext_access_type AccessType) (paddr :: xlenbits) (width :: int) (aq :: bool) (rl :: bool) (res :: bool) (meta :: bool) = ( if ((within_mmio_readable paddr width)) then (mmio_read t paddr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) \ ((\ (w__0 :: (( 'int8_times_n::len)Word.word) MemoryOpResult) . return ((MemoryOpResult_add_meta w__0 default_meta :: ((( 'int8_times_n::len)Word.word * unit)) MemoryOpResult)))) else if ((within_phys_mem paddr width)) then (case ((ext_check_phys_mem_read t paddr width aq rl res meta)) of Ext_PhysAddr_OK (_) => (phys_mem_read t paddr width aq rl res meta :: ( ((( 'int8_times_n::len)Word.word * unit))MemoryOpResult) M) | Ext_PhysAddr_Error (e) => return (MemException e) ) else return ((case t of Execute (_) => MemException (E_Fetch_Access_Fault () ) | Read (Data1) => MemException (E_Load_Access_Fault () ) | _ => MemException (E_SAMO_Access_Fault () ) )))\ for t :: "(ext_access_type)AccessType " and paddr :: "(32)Word.word " and width :: " int " and aq :: " bool " and rl :: " bool " and res :: " bool " and meta :: " bool " \ \\val pmp_mem_read : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> Privilege -> mword ty32 -> integer -> bool -> bool -> bool -> bool -> M (MemoryOpResult ((mword 'int8_times_n * unit)))\\ definition pmp_mem_read :: \(ext_access_type)AccessType \ Privilege \(32)Word.word \ int \ bool \ bool \ bool \ bool \((register_value),((('int8_times_n::len)Word.word*unit)MemoryOpResult),(exception))monad \ where \ pmp_mem_read (t :: ext_access_type AccessType) (p :: Privilege) (paddr :: xlenbits) (width :: int) (aq :: bool) (rl :: bool) (res :: bool) (meta :: bool) = ( if ((\ ((plat_enable_pmp () )))) then (checked_mem_read t paddr width aq rl res meta :: ( ((( 'int8_times_n::len)Word.word * unit))MemoryOpResult) M) else pmpCheck paddr width t p \ ((\ (w__1 :: ExceptionType option) . (case w__1 of None => (checked_mem_read t paddr width aq rl res meta :: ( ((( 'int8_times_n::len)Word.word * unit))MemoryOpResult) M) | Some (e) => return (MemException e) ))))\ for t :: "(ext_access_type)AccessType " and p :: " Privilege " and paddr :: "(32)Word.word " and width :: " int " and aq :: " bool " and rl :: " bool " and res :: " bool " and meta :: " bool " \ \\val rvfi_read : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> MemoryOpResult ((mword 'int8_times_n * unit)) -> unit\\ definition rvfi_read :: \(32)Word.word \ int \(('int8_times_n::len)Word.word*unit)MemoryOpResult \ unit \ where \ rvfi_read addr width result = ( () )\ for addr :: "(32)Word.word " and width :: " int " and result :: "(('int8_times_n::len)Word.word*unit)MemoryOpResult " \ \\val mem_read : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> mword ty32 -> integer -> bool -> bool -> bool -> M (MemoryOpResult (mword 'int8_times_n))\\ \ \\val mem_read_priv : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> Privilege -> mword ty32 -> integer -> bool -> bool -> bool -> M (MemoryOpResult (mword 'int8_times_n))\\ \ \\val mem_read_meta : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> mword ty32 -> integer -> bool -> bool -> bool -> bool -> M (MemoryOpResult ((mword 'int8_times_n * unit)))\\ \ \\val mem_read_priv_meta : forall 'int8_times_n. Size 'int8_times_n => AccessType unit -> Privilege -> mword ty32 -> integer -> bool -> bool -> bool -> bool -> M (MemoryOpResult ((mword 'int8_times_n * unit)))\\ definition mem_read_priv_meta :: \(unit)AccessType \ Privilege \(32)Word.word \ int \ bool \ bool \ bool \ bool \((register_value),((('int8_times_n::len)Word.word*unit)MemoryOpResult),(exception))monad \ where \ mem_read_priv_meta typ1 priv paddr width aq rl res meta = ( (if ((((((aq \ res))) \ ((\ ((is_aligned_addr paddr width))))))) then return (MemException (E_Load_Addr_Align () )) else (case (aq, rl, res) of (False, True, False) => throw (Error_not_implemented (''load.rl'')) | (False, True, True) => throw (Error_not_implemented (''lr.rl'')) | (_, _, _) => (pmp_mem_read typ1 priv paddr width aq rl res meta :: ( ((( 'int8_times_n::len)Word.word * unit))MemoryOpResult) M) )) \ ((\ (result :: ((( 'int8_times_n::len)Word.word * unit)) MemoryOpResult) . (let (_ :: unit) = (rvfi_read paddr width result) in return result))))\ for typ1 :: "(unit)AccessType " and priv :: " Privilege " and paddr :: "(32)Word.word " and width :: " int " and aq :: " bool " and rl :: " bool " and res :: " bool " and meta :: " bool " definition mem_read_meta :: \(unit)AccessType \(32)Word.word \ int \ bool \ bool \ bool \ bool \((register_value),((('int8_times_n::len)Word.word*unit)MemoryOpResult),(exception))monad \ where \ mem_read_meta typ1 paddr width aq rl res meta = ( read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . read_reg cur_privilege_ref \ ((\ (w__1 :: Privilege) . effectivePrivilege typ1 w__0 w__1 \ ((\ (w__2 :: Privilege) . (mem_read_priv_meta typ1 w__2 paddr width aq rl res meta :: ( ((( 'int8_times_n::len)Word.word * unit))MemoryOpResult) M))))))))\ for typ1 :: "(unit)AccessType " and paddr :: "(32)Word.word " and width :: " int " and aq :: " bool " and rl :: " bool " and res :: " bool " and meta :: " bool " definition mem_read_priv :: \(unit)AccessType \ Privilege \(32)Word.word \ int \ bool \ bool \ bool \((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad \ where \ mem_read_priv typ1 priv paddr width aq rl res = ( (mem_read_priv_meta typ1 priv paddr width aq rl res False :: ( ((( 'int8_times_n::len)Word.word * unit))MemoryOpResult) M) \ ((\ (w__0 :: ((( 'int8_times_n::len)Word.word * unit)) MemoryOpResult) . return ((MemoryOpResult_drop_meta w__0 :: (( 'int8_times_n::len)Word.word) MemoryOpResult)))))\ for typ1 :: "(unit)AccessType " and priv :: " Privilege " and paddr :: "(32)Word.word " and width :: " int " and aq :: " bool " and rl :: " bool " and res :: " bool " definition mem_read :: \(unit)AccessType \(32)Word.word \ int \ bool \ bool \ bool \((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad \ where \ mem_read typ1 paddr width aq rel res = ( read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . read_reg cur_privilege_ref \ ((\ (w__1 :: Privilege) . effectivePrivilege typ1 w__0 w__1 \ ((\ (w__2 :: Privilege) . (mem_read_priv typ1 w__2 paddr width aq rel res :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M))))))))\ for typ1 :: "(unit)AccessType " and paddr :: "(32)Word.word " and width :: " int " and aq :: " bool " and rel :: " bool " and res :: " bool " \ \\val mem_write_ea : mword ty32 -> integer -> bool -> bool -> bool -> M (MemoryOpResult unit)\\ definition mem_write_ea :: \(32)Word.word \ int \ bool \ bool \ bool \((register_value),((unit)MemoryOpResult),(exception))monad \ where \ mem_write_ea addr width aq rl con = ( if ((((((rl \ con))) \ ((\ ((is_aligned_addr addr width))))))) then return (MemException (E_SAMO_Addr_Align () )) else (case (aq, rl, con) of (False, False, False) => write_ram_ea Write_plain addr width \ return (MemValue () ) | (False, True, False) => write_ram_ea Write_RISCV_release addr width \ return (MemValue () ) | (False, False, True) => write_ram_ea Write_RISCV_conditional addr width \ return (MemValue () ) | (False, True, True) => write_ram_ea Write_RISCV_conditional_release addr width \ return (MemValue () ) | (True, False, False) => throw (Error_not_implemented (''store.aq'')) | (True, True, False) => write_ram_ea Write_RISCV_strong_release addr width \ return (MemValue () ) | (True, False, True) => throw (Error_not_implemented (''sc.aq'')) | (True, True, True) => write_ram_ea Write_RISCV_conditional_strong_release addr width \ return (MemValue () ) ))\ for addr :: "(32)Word.word " and width :: " int " and aq :: " bool " and rl :: " bool " and con :: " bool " \ \\val rvfi_write : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> mword 'int8_times_n -> unit -> unit\\ definition rvfi_write :: \(32)Word.word \ int \('int8_times_n::len)Word.word \ unit \ unit \ where \ rvfi_write addr width value1 meta = ( () )\ for addr :: "(32)Word.word " and width :: " int " and value1 :: "('int8_times_n::len)Word.word " and meta :: " unit " \ \\val phys_mem_write : forall 'int8_times_n. Size 'int8_times_n => write_kind -> mword ty32 -> integer -> mword 'int8_times_n -> unit -> M (MemoryOpResult bool)\\ definition phys_mem_write :: \ write_kind \(32)Word.word \ int \('int8_times_n::len)Word.word \ unit \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ phys_mem_write (wk :: write_kind) (paddr :: xlenbits) (width :: int) (data :: ( 'int8_times_n::len)Word.word) (meta :: mem_meta) = ( (let (_ :: unit) = (rvfi_write paddr width data meta) in write_ram wk paddr width data meta \ ((\ (w__0 :: bool) . (let result = (MemValue w__0) in (let (_ :: unit) = (if ((get_config_print_mem () )) then print_dbg (((@) (''mem['') (((@) ((string_of_bits paddr)) (((@) (''] <- '') ((string_of_bits data)))))))) else () ) in return result))))))\ for wk :: " write_kind " and paddr :: "(32)Word.word " and width :: " int " and data :: "('int8_times_n::len)Word.word " and meta :: " unit " \ \\val checked_mem_write : forall 'int8_times_n. Size 'int8_times_n => write_kind -> mword ty32 -> integer -> mword 'int8_times_n -> unit -> M (MemoryOpResult bool)\\ definition checked_mem_write :: \ write_kind \(32)Word.word \ int \('int8_times_n::len)Word.word \ unit \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ checked_mem_write (wk :: write_kind) (paddr :: xlenbits) (width :: int) (data :: ( 'int8_times_n::len)Word.word) (meta :: mem_meta) = ( if ((within_mmio_writable paddr width)) then mmio_write paddr width data else if ((within_phys_mem paddr width)) then (case ((ext_check_phys_mem_write wk paddr width data meta)) of Ext_PhysAddr_OK (_) => phys_mem_write wk paddr width data meta | Ext_PhysAddr_Error (e) => return (MemException e) ) else return (MemException (E_SAMO_Access_Fault () )))\ for wk :: " write_kind " and paddr :: "(32)Word.word " and width :: " int " and data :: "('int8_times_n::len)Word.word " and meta :: " unit " \ \\val pmp_mem_write : forall 'int8_times_n. Size 'int8_times_n => write_kind -> mword ty32 -> integer -> mword 'int8_times_n -> AccessType unit -> Privilege -> unit -> M (MemoryOpResult bool)\\ definition pmp_mem_write :: \ write_kind \(32)Word.word \ int \('int8_times_n::len)Word.word \(ext_access_type)AccessType \ Privilege \ unit \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ pmp_mem_write (wk :: write_kind) (paddr :: xlenbits) (width :: int) (data :: ( 'int8_times_n::len)Word.word) (typ1 :: ext_access_type AccessType) (priv :: Privilege) (meta :: mem_meta) = ( if ((\ ((plat_enable_pmp () )))) then checked_mem_write wk paddr width data meta else pmpCheck paddr width typ1 priv \ ((\ (w__1 :: ExceptionType option) . (case w__1 of None => checked_mem_write wk paddr width data meta | Some (e) => return (MemException e) ))))\ for wk :: " write_kind " and paddr :: "(32)Word.word " and width :: " int " and data :: "('int8_times_n::len)Word.word " and typ1 :: "(ext_access_type)AccessType " and priv :: " Privilege " and meta :: " unit " \ \\val mem_write_value_priv_meta : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> mword 'int8_times_n -> AccessType unit -> Privilege -> unit -> bool -> bool -> bool -> M (MemoryOpResult bool)\\ definition mem_write_value_priv_meta :: \(32)Word.word \ int \('int8_times_n::len)Word.word \(unit)AccessType \ Privilege \ unit \ bool \ bool \ bool \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ mem_write_value_priv_meta paddr width value1 typ1 priv meta aq rl con = ( (let (_ :: unit) = (rvfi_write paddr width value1 meta) in if ((((((rl \ con))) \ ((\ ((is_aligned_addr paddr width))))))) then return (MemException (E_SAMO_Addr_Align () )) else (case (aq, rl, con) of (False, False, False) => return Write_plain | (False, True, False) => return Write_RISCV_release | (False, False, True) => return Write_RISCV_conditional | (False, True, True) => return Write_RISCV_conditional_release | (True, True, False) => return Write_RISCV_strong_release | (True, True, True) => return Write_RISCV_conditional_strong_release | (True, False, False) => throw (Error_not_implemented (''store.aq'')) | (True, False, True) => throw (Error_not_implemented (''sc.aq'')) ) \ ((\ (wk :: write_kind) . pmp_mem_write wk paddr width value1 typ1 priv meta))))\ for paddr :: "(32)Word.word " and width :: " int " and value1 :: "('int8_times_n::len)Word.word " and typ1 :: "(unit)AccessType " and priv :: " Privilege " and meta :: " unit " and aq :: " bool " and rl :: " bool " and con :: " bool " \ \\val mem_write_value_priv : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> mword 'int8_times_n -> Privilege -> bool -> bool -> bool -> M (MemoryOpResult bool)\\ definition mem_write_value_priv :: \(32)Word.word \ int \('int8_times_n::len)Word.word \ Privilege \ bool \ bool \ bool \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ mem_write_value_priv paddr width value1 priv aq rl con = ( mem_write_value_priv_meta paddr width value1 (Write default_write_acc) priv default_meta aq rl con )\ for paddr :: "(32)Word.word " and width :: " int " and value1 :: "('int8_times_n::len)Word.word " and priv :: " Privilege " and aq :: " bool " and rl :: " bool " and con :: " bool " \ \\val mem_write_value_meta : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> mword 'int8_times_n -> unit -> unit -> bool -> bool -> bool -> M (MemoryOpResult bool)\\ definition mem_write_value_meta :: \(32)Word.word \ int \('int8_times_n::len)Word.word \ unit \ unit \ bool \ bool \ bool \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ mem_write_value_meta paddr width value1 ext_acc meta aq rl con = ( (let typ1 = (Write ext_acc) in read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . read_reg cur_privilege_ref \ ((\ (w__1 :: Privilege) . effectivePrivilege typ1 w__0 w__1 \ ((\ ep . mem_write_value_priv_meta paddr width value1 typ1 ep meta aq rl con))))))))\ for paddr :: "(32)Word.word " and width :: " int " and value1 :: "('int8_times_n::len)Word.word " and ext_acc :: " unit " and meta :: " unit " and aq :: " bool " and rl :: " bool " and con :: " bool " \ \\val mem_write_value : forall 'int8_times_n. Size 'int8_times_n => mword ty32 -> integer -> mword 'int8_times_n -> bool -> bool -> bool -> M (MemoryOpResult bool)\\ definition mem_write_value :: \(32)Word.word \ int \('int8_times_n::len)Word.word \ bool \ bool \ bool \((register_value),((bool)MemoryOpResult),(exception))monad \ where \ mem_write_value paddr width value1 aq rl con = ( mem_write_value_meta paddr width value1 default_write_acc default_meta aq rl con )\ for paddr :: "(32)Word.word " and width :: " int " and value1 :: "('int8_times_n::len)Word.word " and aq :: " bool " and rl :: " bool " and con :: " bool " definition default_sv32_ext_pte :: \(10)Word.word \ where \ default_sv32_ext_pte = ( ( 0b0000000000 :: 10 Word.word))\ \ \\val undefined_PTE_Bits : unit -> M PTE_Bits\\ definition undefined_PTE_Bits :: \ unit \((register_value),(PTE_Bits),(exception))monad \ where \ undefined_PTE_Bits _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 8 Word.word) M) \ ((\ (w__0 :: 8 Word.word) . return ((| PTE_Bits_bits = w__0 |)))))\ \ \\val Mk_PTE_Bits : mword ty8 -> PTE_Bits\\ definition Mk_PTE_Bits :: \(8)Word.word \ PTE_Bits \ where \ Mk_PTE_Bits v = ( (| PTE_Bits_bits = v |) )\ for v :: "(8)Word.word " definition get_PTE_Bits_bits :: \ PTE_Bits \(8)Word.word \ where \ get_PTE_Bits_bits v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_bits :: \((regstate),(register_value),(PTE_Bits))register_ref \(8)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 7 :: int)::ii) (( 0 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(8)Word.word " definition update_PTE_Bits_bits :: \ PTE_Bits \(8)Word.word \ PTE_Bits \ where \ update_PTE_Bits_bits v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(8)Word.word " definition get_PTE_Bits_A :: \ PTE_Bits \(1)Word.word \ where \ get_PTE_Bits_A v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_A :: \((regstate),(register_value),(PTE_Bits))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_A r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 6 :: int)::ii) (( 6 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(1)Word.word " definition update_PTE_Bits_A :: \ PTE_Bits \(1)Word.word \ PTE_Bits \ where \ update_PTE_Bits_A v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 6 :: int)::ii) (( 6 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(1)Word.word " definition get_PTE_Bits_D :: \ PTE_Bits \(1)Word.word \ where \ get_PTE_Bits_D v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_D :: \((regstate),(register_value),(PTE_Bits))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_D r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 7 :: int)::ii) (( 7 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(1)Word.word " definition update_PTE_Bits_D :: \ PTE_Bits \(1)Word.word \ PTE_Bits \ where \ update_PTE_Bits_D v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 7 :: int)::ii) (( 7 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(1)Word.word " definition get_PTE_Bits_G :: \ PTE_Bits \(1)Word.word \ where \ get_PTE_Bits_G v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_G :: \((regstate),(register_value),(PTE_Bits))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_G r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 5 :: int)::ii) (( 5 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(1)Word.word " definition update_PTE_Bits_G :: \ PTE_Bits \(1)Word.word \ PTE_Bits \ where \ update_PTE_Bits_G v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 5 :: int)::ii) (( 5 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(1)Word.word " definition get_PTE_Bits_R :: \ PTE_Bits \(1)Word.word \ where \ get_PTE_Bits_R v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_R :: \((regstate),(register_value),(PTE_Bits))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_R r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 1 :: int)::ii) (( 1 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(1)Word.word " definition update_PTE_Bits_R :: \ PTE_Bits \(1)Word.word \ PTE_Bits \ where \ update_PTE_Bits_R v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 1 :: int)::ii) (( 1 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(1)Word.word " definition get_PTE_Bits_U :: \ PTE_Bits \(1)Word.word \ where \ get_PTE_Bits_U v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_U :: \((regstate),(register_value),(PTE_Bits))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_U r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 4 :: int)::ii) (( 4 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(1)Word.word " definition update_PTE_Bits_U :: \ PTE_Bits \(1)Word.word \ PTE_Bits \ where \ update_PTE_Bits_U v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 4 :: int)::ii) (( 4 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(1)Word.word " definition get_PTE_Bits_V :: \ PTE_Bits \(1)Word.word \ where \ get_PTE_Bits_V v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_V :: \((regstate),(register_value),(PTE_Bits))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_V r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 0 :: int)::ii) (( 0 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(1)Word.word " definition update_PTE_Bits_V :: \ PTE_Bits \(1)Word.word \ PTE_Bits \ where \ update_PTE_Bits_V v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 0 :: int)::ii) (( 0 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(1)Word.word " definition get_PTE_Bits_W :: \ PTE_Bits \(1)Word.word \ where \ get_PTE_Bits_W v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_W :: \((regstate),(register_value),(PTE_Bits))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_W r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 2 :: int)::ii) (( 2 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(1)Word.word " definition update_PTE_Bits_W :: \ PTE_Bits \(1)Word.word \ PTE_Bits \ where \ update_PTE_Bits_W v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 2 :: int)::ii) (( 2 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(1)Word.word " definition get_PTE_Bits_X :: \ PTE_Bits \(1)Word.word \ where \ get_PTE_Bits_X v = ( (subrange_vec_dec(PTE_Bits_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word))\ for v :: " PTE_Bits " definition set_PTE_Bits_X :: \((regstate),(register_value),(PTE_Bits))register_ref \(1)Word.word \((register_value),(unit),(exception))monad \ where \ set_PTE_Bits_X r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits r) (( 3 :: int)::ii) (( 3 :: int)::ii) v :: 8 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(PTE_Bits))register_ref " and v :: "(1)Word.word " definition update_PTE_Bits_X :: \ PTE_Bits \(1)Word.word \ PTE_Bits \ where \ update_PTE_Bits_X v x = ( ( v (| PTE_Bits_bits := ((update_subrange_vec_dec(PTE_Bits_bits v) (( 3 :: int)::ii) (( 3 :: int)::ii) x :: 8 Word.word)) |)))\ for v :: " PTE_Bits " and x :: "(1)Word.word " \ \\val isPTEPtr : mword ty8 -> mword ty10 -> bool\\ definition isPTEPtr :: \(8)Word.word \(10)Word.word \ bool \ where \ isPTEPtr (p :: pteAttribs) (ext :: extPte) = ( (let a = (Mk_PTE_Bits p) in ((((((get_PTE_Bits_R a :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_W a :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((get_PTE_Bits_X a :: 1 Word.word)) = ( 0b0 :: 1 Word.word))))))))))\ for p :: "(8)Word.word " and ext :: "(10)Word.word " \ \\val isInvalidPTE : mword ty8 -> mword ty10 -> bool\\ definition isInvalidPTE :: \(8)Word.word \(10)Word.word \ bool \ where \ isInvalidPTE (p :: pteAttribs) (ext :: extPte) = ( (let a = (Mk_PTE_Bits p) in ((((((get_PTE_Bits_V a :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_W a :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ (((((get_PTE_Bits_R a :: 1 Word.word)) = ( 0b0 :: 1 Word.word))))))))))\ for p :: "(8)Word.word " and ext :: "(10)Word.word " \ \\val to_pte_check : bool -> PTE_Check\\ definition to_pte_check :: \ bool \ PTE_Check \ where \ to_pte_check b = ( if b then PTE_Check_Success () else PTE_Check_Failure (() , () ))\ for b :: " bool " \ \\val checkPTEPermission : AccessType unit -> Privilege -> bool -> bool -> PTE_Bits -> mword ty10 -> unit -> M PTE_Check\\ fun checkPTEPermission :: \(ext_access_type)AccessType \ Privilege \ bool \ bool \ PTE_Bits \(10)Word.word \ unit \((register_value),(PTE_Check),(exception))monad \ where \ checkPTEPermission ((Read (_)) :: ext_access_type AccessType) (User :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( return ((to_pte_check ((((((((get_PTE_Bits_U p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_R p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_X p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ mxr))))))))))))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " |\ checkPTEPermission ((Write (_)) :: ext_access_type AccessType) (User :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( return ((to_pte_check ((((((((get_PTE_Bits_U p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ (((((get_PTE_Bits_W p :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))))))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " |\ checkPTEPermission ((ReadWrite ((_, _))) :: ext_access_type AccessType) (User :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( return ((to_pte_check ((((((((get_PTE_Bits_U p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_W p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_R p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_X p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ mxr)))))))))))))))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " |\ checkPTEPermission ((Execute (_)) :: ext_access_type AccessType) (User :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( return ((to_pte_check ((((((((get_PTE_Bits_U p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ (((((get_PTE_Bits_X p :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))))))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " |\ checkPTEPermission ((Read (_)) :: ext_access_type AccessType) (Supervisor :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( return ((to_pte_check (((((((((((get_PTE_Bits_U p :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ do_sum))) \ ((((((((get_PTE_Bits_R p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_X p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ mxr))))))))))))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " |\ checkPTEPermission ((Write (_)) :: ext_access_type AccessType) (Supervisor :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( return ((to_pte_check (((((((((((get_PTE_Bits_U p :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ do_sum))) \ (((((get_PTE_Bits_W p :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))))))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " |\ checkPTEPermission ((ReadWrite ((_, _))) :: ext_access_type AccessType) (Supervisor :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( return ((to_pte_check (((((((((((get_PTE_Bits_U p :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ do_sum))) \ ((((((((get_PTE_Bits_W p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_R p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ ((((((((get_PTE_Bits_X p :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) \ mxr)))))))))))))))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " |\ checkPTEPermission ((Execute (_)) :: ext_access_type AccessType) (Supervisor :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( return ((to_pte_check ((((((((get_PTE_Bits_U p :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((get_PTE_Bits_X p :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))))))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " |\ checkPTEPermission (_ :: ext_access_type AccessType) (Machine :: Privilege) (mxr :: bool) (do_sum :: bool) (p :: PTE_Bits) (ext :: extPte) (ext_ptw :: ext_ptw) = ( internal_error (''m-mode mem perm check''))\ for mxr :: " bool " and do_sum :: " bool " and p :: " PTE_Bits " and ext :: "(10)Word.word " and ext_ptw :: " unit " \ \\val update_PTE_Bits : PTE_Bits -> AccessType unit -> mword ty10 -> maybe ((PTE_Bits * mword ty10))\\ definition update_PTE_Bits :: \ PTE_Bits \(ext_access_type)AccessType \(10)Word.word \(PTE_Bits*(10)Word.word)option \ where \ update_PTE_Bits (p :: PTE_Bits) (a :: ext_access_type AccessType) (ext :: extPte) = ( (let update_d = ((((((get_PTE_Bits_D p :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ ((case a of Execute (_) => False | Read (_) => False | Write (_) => True | ReadWrite ((_, _)) => True ))) in (let update_a = (((get_PTE_Bits_A p :: 1 Word.word)) = ( 0b0 :: 1 Word.word)) in if (((update_d \ update_a))) then (let np = (update_PTE_Bits_A p ( 0b1 :: 1 Word.word)) in (let np = (if update_d then update_PTE_Bits_D np ( 0b1 :: 1 Word.word) else np) in Some (np, ext))) else None)))\ for p :: " PTE_Bits " and a :: "(ext_access_type)AccessType " and ext :: "(10)Word.word " fun ptw_error_to_str :: \ PTW_Error \ string \ where \ ptw_error_to_str (PTW_Invalid_Addr (_)) = ( (''invalid-source-addr''))\ |\ ptw_error_to_str (PTW_Access (_)) = ( (''mem-access-error''))\ |\ ptw_error_to_str (PTW_Invalid_PTE (_)) = ( (''invalid-pte''))\ |\ ptw_error_to_str (PTW_No_Permission (_)) = ( (''no-permission''))\ |\ ptw_error_to_str (PTW_Misaligned (_)) = ( (''misaligned-superpage''))\ |\ ptw_error_to_str (PTW_PTE_Update (_)) = ( (''pte-update-needed''))\ |\ ptw_error_to_str (PTW_Ext_Error (e)) = ( (''extension-error''))\ for e :: " unit " \ \\val ext_get_ptw_error : unit -> PTW_Error\\ definition ext_get_ptw_error :: \ unit \ PTW_Error \ where \ ext_get_ptw_error eptwf = ( PTW_No_Permission () )\ for eptwf :: " unit " \ \\val translationException : AccessType unit -> PTW_Error -> ExceptionType\\ fun translationException :: \(ext_access_type)AccessType \ PTW_Error \ ExceptionType \ where \ translationException (_ :: ext_access_type AccessType) ((PTW_Ext_Error (e)) :: PTW_Error) = ( E_Extension ((ext_translate_exception e)))\ for e :: " unit " |\ translationException ((ReadWrite (_)) :: ext_access_type AccessType) ((PTW_Access (_)) :: PTW_Error) = ( E_SAMO_Access_Fault () )\ |\ translationException ((ReadWrite (_)) :: ext_access_type AccessType) (_ :: PTW_Error) = ( E_SAMO_Page_Fault () )\ |\ translationException ((Read (_)) :: ext_access_type AccessType) ((PTW_Access (_)) :: PTW_Error) = ( E_Load_Access_Fault () )\ |\ translationException ((Read (_)) :: ext_access_type AccessType) (_ :: PTW_Error) = ( E_Load_Page_Fault () )\ |\ translationException ((Write (_)) :: ext_access_type AccessType) ((PTW_Access (_)) :: PTW_Error) = ( E_SAMO_Access_Fault () )\ |\ translationException ((Write (_)) :: ext_access_type AccessType) (_ :: PTW_Error) = ( E_SAMO_Page_Fault () )\ |\ translationException ((Execute (_)) :: ext_access_type AccessType) ((PTW_Access (_)) :: PTW_Error) = ( E_Fetch_Access_Fault () )\ |\ translationException ((Execute (_)) :: ext_access_type AccessType) (_ :: PTW_Error) = ( E_Fetch_Page_Fault () )\ definition PAGESIZE_BITS :: \ int \ where \ PAGESIZE_BITS = ( (( 12 :: int)::ii))\ \ \\val curAsid32 : mword ty32 -> mword ty9\\ definition curAsid32 :: \(32)Word.word \(9)Word.word \ where \ curAsid32 satp1 = ( (let s = (Mk_Satp32 satp1) in (get_Satp32_Asid s :: 9 Word.word)))\ for satp1 :: "(32)Word.word " \ \\val curPTB32 : mword ty32 -> mword ty34\\ definition curPTB32 :: \(32)Word.word \(34)Word.word \ where \ curPTB32 satp1 = ( (let (s :: Satp32) = (Mk_Satp32 satp1) in (shiftl ((EXTZ (( 34 :: int)::ii) ((get_Satp32_PPN s :: 22 Word.word)) :: 34 Word.word)) PAGESIZE_BITS :: 34 Word.word)))\ for satp1 :: "(32)Word.word " definition SV32_LEVEL_BITS :: \ int \ where \ SV32_LEVEL_BITS = ( (( 10 :: int)::ii))\ definition SV32_LEVELS :: \ int \ where \ SV32_LEVELS = ( (( 2 :: int)::ii))\ definition PTE32_LOG_SIZE :: \ int \ where \ PTE32_LOG_SIZE = ( (( 2 :: int)::ii))\ definition PTE32_SIZE :: \ int \ where \ PTE32_SIZE = ( (( 4 :: int)::ii))\ \ \\val undefined_SV32_Vaddr : unit -> M SV32_Vaddr\\ definition undefined_SV32_Vaddr :: \ unit \((register_value),(SV32_Vaddr),(exception))monad \ where \ undefined_SV32_Vaddr _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| SV32_Vaddr_bits = w__0 |)))))\ \ \\val Mk_SV32_Vaddr : mword ty32 -> SV32_Vaddr\\ definition Mk_SV32_Vaddr :: \(32)Word.word \ SV32_Vaddr \ where \ Mk_SV32_Vaddr v = ( (| SV32_Vaddr_bits = v |) )\ for v :: "(32)Word.word " definition get_SV32_Vaddr_bits :: \ SV32_Vaddr \(32)Word.word \ where \ get_SV32_Vaddr_bits v = ( (subrange_vec_dec(SV32_Vaddr_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " SV32_Vaddr " definition set_SV32_Vaddr_bits :: \((regstate),(register_value),(SV32_Vaddr))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_Vaddr_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_Vaddr_bits := ((update_subrange_vec_dec(SV32_Vaddr_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_Vaddr))register_ref " and v :: "(32)Word.word " definition update_SV32_Vaddr_bits :: \ SV32_Vaddr \(32)Word.word \ SV32_Vaddr \ where \ update_SV32_Vaddr_bits v x = ( ( v (| SV32_Vaddr_bits := ((update_subrange_vec_dec(SV32_Vaddr_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " SV32_Vaddr " and x :: "(32)Word.word " \ \\val _get_SV32_Vaddr_PgOfs : SV32_Vaddr -> mword ty12\\ definition get_SV32_Vaddr_PgOfs :: \ SV32_Vaddr \(12)Word.word \ where \ get_SV32_Vaddr_PgOfs v = ( (subrange_vec_dec(SV32_Vaddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) :: 12 Word.word))\ for v :: " SV32_Vaddr " \ \\val _set_SV32_Vaddr_PgOfs : register_ref regstate register_value SV32_Vaddr -> mword ty12 -> M unit\\ definition set_SV32_Vaddr_PgOfs :: \((regstate),(register_value),(SV32_Vaddr))register_ref \(12)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_Vaddr_PgOfs r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_Vaddr_bits := ((update_subrange_vec_dec(SV32_Vaddr_bits r) (( 11 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_Vaddr))register_ref " and v :: "(12)Word.word " \ \\val _update_SV32_Vaddr_PgOfs : SV32_Vaddr -> mword ty12 -> SV32_Vaddr\\ definition update_SV32_Vaddr_PgOfs :: \ SV32_Vaddr \(12)Word.word \ SV32_Vaddr \ where \ update_SV32_Vaddr_PgOfs v x = ( ( v (| SV32_Vaddr_bits := ((update_subrange_vec_dec(SV32_Vaddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " SV32_Vaddr " and x :: "(12)Word.word " \ \\val _update_SV32_Paddr_PgOfs : SV32_Paddr -> mword ty12 -> SV32_Paddr\\ \ \\val _update_SV39_Paddr_PgOfs : SV39_Paddr -> mword ty12 -> SV39_Paddr\\ \ \\val _update_SV39_Vaddr_PgOfs : SV39_Vaddr -> mword ty12 -> SV39_Vaddr\\ \ \\val _update_SV48_Paddr_PgOfs : SV48_Paddr -> mword ty12 -> SV48_Paddr\\ \ \\val _update_SV48_Vaddr_PgOfs : SV48_Vaddr -> mword ty12 -> SV48_Vaddr\\ \ \\val _get_SV32_Paddr_PgOfs : SV32_Paddr -> mword ty12\\ \ \\val _get_SV39_Paddr_PgOfs : SV39_Paddr -> mword ty12\\ \ \\val _get_SV39_Vaddr_PgOfs : SV39_Vaddr -> mword ty12\\ \ \\val _get_SV48_Paddr_PgOfs : SV48_Paddr -> mword ty12\\ \ \\val _get_SV48_Vaddr_PgOfs : SV48_Vaddr -> mword ty12\\ \ \\val _set_SV32_Paddr_PgOfs : register_ref regstate register_value SV32_Paddr -> mword ty12 -> M unit\\ \ \\val _set_SV39_Paddr_PgOfs : register_ref regstate register_value SV39_Paddr -> mword ty12 -> M unit\\ \ \\val _set_SV39_Vaddr_PgOfs : register_ref regstate register_value SV39_Vaddr -> mword ty12 -> M unit\\ \ \\val _set_SV48_Paddr_PgOfs : register_ref regstate register_value SV48_Paddr -> mword ty12 -> M unit\\ \ \\val _set_SV48_Vaddr_PgOfs : register_ref regstate register_value SV48_Vaddr -> mword ty12 -> M unit\\ \ \\val _get_SV32_Vaddr_VPNi : SV32_Vaddr -> mword ty20\\ definition get_SV32_Vaddr_VPNi :: \ SV32_Vaddr \(20)Word.word \ where \ get_SV32_Vaddr_VPNi v = ( (subrange_vec_dec(SV32_Vaddr_bits v) (( 31 :: int)::ii) (( 12 :: int)::ii) :: 20 Word.word))\ for v :: " SV32_Vaddr " \ \\val _set_SV32_Vaddr_VPNi : register_ref regstate register_value SV32_Vaddr -> mword ty20 -> M unit\\ definition set_SV32_Vaddr_VPNi :: \((regstate),(register_value),(SV32_Vaddr))register_ref \(20)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_Vaddr_VPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_Vaddr_bits := ((update_subrange_vec_dec(SV32_Vaddr_bits r) (( 31 :: int)::ii) (( 12 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_Vaddr))register_ref " and v :: "(20)Word.word " \ \\val _update_SV32_Vaddr_VPNi : SV32_Vaddr -> mword ty20 -> SV32_Vaddr\\ definition update_SV32_Vaddr_VPNi :: \ SV32_Vaddr \(20)Word.word \ SV32_Vaddr \ where \ update_SV32_Vaddr_VPNi v x = ( ( v (| SV32_Vaddr_bits := ((update_subrange_vec_dec(SV32_Vaddr_bits v) (( 31 :: int)::ii) (( 12 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " SV32_Vaddr " and x :: "(20)Word.word " \ \\val _update_SV39_Vaddr_VPNi : SV39_Vaddr -> mword ty27 -> SV39_Vaddr\\ \ \\val _update_SV48_Vaddr_VPNi : SV48_Vaddr -> mword ty27 -> SV48_Vaddr\\ \ \\val _get_SV39_Vaddr_VPNi : SV39_Vaddr -> mword ty27\\ \ \\val _get_SV48_Vaddr_VPNi : SV48_Vaddr -> mword ty27\\ \ \\val _set_SV39_Vaddr_VPNi : register_ref regstate register_value SV39_Vaddr -> mword ty27 -> M unit\\ \ \\val _set_SV48_Vaddr_VPNi : register_ref regstate register_value SV48_Vaddr -> mword ty27 -> M unit\\ \ \\val undefined_SV32_Paddr : unit -> M SV32_Paddr\\ definition undefined_SV32_Paddr :: \ unit \((register_value),(SV32_Paddr),(exception))monad \ where \ undefined_SV32_Paddr _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 34 Word.word) M) \ ((\ (w__0 :: 34 Word.word) . return ((| SV32_Paddr_bits = w__0 |)))))\ \ \\val Mk_SV32_Paddr : mword ty34 -> SV32_Paddr\\ definition Mk_SV32_Paddr :: \(34)Word.word \ SV32_Paddr \ where \ Mk_SV32_Paddr v = ( (| SV32_Paddr_bits = v |) )\ for v :: "(34)Word.word " definition get_SV32_Paddr_bits :: \ SV32_Paddr \(34)Word.word \ where \ get_SV32_Paddr_bits v = ( (subrange_vec_dec(SV32_Paddr_bits v) (( 33 :: int)::ii) (( 0 :: int)::ii) :: 34 Word.word))\ for v :: " SV32_Paddr " definition set_SV32_Paddr_bits :: \((regstate),(register_value),(SV32_Paddr))register_ref \(34)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_Paddr_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_Paddr_bits := ((update_subrange_vec_dec(SV32_Paddr_bits r) (( 33 :: int)::ii) (( 0 :: int)::ii) v :: 34 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_Paddr))register_ref " and v :: "(34)Word.word " definition update_SV32_Paddr_bits :: \ SV32_Paddr \(34)Word.word \ SV32_Paddr \ where \ update_SV32_Paddr_bits v x = ( ( v (| SV32_Paddr_bits := ((update_subrange_vec_dec(SV32_Paddr_bits v) (( 33 :: int)::ii) (( 0 :: int)::ii) x :: 34 Word.word)) |)))\ for v :: " SV32_Paddr " and x :: "(34)Word.word " \ \\val _get_SV32_Paddr_PPNi : SV32_Paddr -> mword ty22\\ definition get_SV32_Paddr_PPNi :: \ SV32_Paddr \(22)Word.word \ where \ get_SV32_Paddr_PPNi v = ( (subrange_vec_dec(SV32_Paddr_bits v) (( 33 :: int)::ii) (( 12 :: int)::ii) :: 22 Word.word))\ for v :: " SV32_Paddr " \ \\val _set_SV32_Paddr_PPNi : register_ref regstate register_value SV32_Paddr -> mword ty22 -> M unit\\ definition set_SV32_Paddr_PPNi :: \((regstate),(register_value),(SV32_Paddr))register_ref \(22)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_Paddr_PPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_Paddr_bits := ((update_subrange_vec_dec(SV32_Paddr_bits r) (( 33 :: int)::ii) (( 12 :: int)::ii) v :: 34 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_Paddr))register_ref " and v :: "(22)Word.word " \ \\val _update_SV32_Paddr_PPNi : SV32_Paddr -> mword ty22 -> SV32_Paddr\\ definition update_SV32_Paddr_PPNi :: \ SV32_Paddr \(22)Word.word \ SV32_Paddr \ where \ update_SV32_Paddr_PPNi v x = ( ( v (| SV32_Paddr_bits := ((update_subrange_vec_dec(SV32_Paddr_bits v) (( 33 :: int)::ii) (( 12 :: int)::ii) x :: 34 Word.word)) |)))\ for v :: " SV32_Paddr " and x :: "(22)Word.word " \ \\val _update_SV32_PTE_PPNi : SV32_PTE -> mword ty22 -> SV32_PTE\\ \ \\val _update_SV39_PTE_PPNi : SV39_PTE -> mword ty44 -> SV39_PTE\\ \ \\val _update_SV39_Paddr_PPNi : SV39_Paddr -> mword ty44 -> SV39_Paddr\\ \ \\val _update_SV48_PTE_PPNi : SV48_PTE -> mword ty44 -> SV48_PTE\\ \ \\val _update_SV48_Paddr_PPNi : SV48_Paddr -> mword ty44 -> SV48_Paddr\\ \ \\val _get_SV32_PTE_PPNi : SV32_PTE -> mword ty22\\ \ \\val _get_SV39_PTE_PPNi : SV39_PTE -> mword ty44\\ \ \\val _get_SV39_Paddr_PPNi : SV39_Paddr -> mword ty44\\ \ \\val _get_SV48_PTE_PPNi : SV48_PTE -> mword ty44\\ \ \\val _get_SV48_Paddr_PPNi : SV48_Paddr -> mword ty44\\ \ \\val _set_SV32_PTE_PPNi : register_ref regstate register_value SV32_PTE -> mword ty22 -> M unit\\ \ \\val _set_SV39_PTE_PPNi : register_ref regstate register_value SV39_PTE -> mword ty44 -> M unit\\ \ \\val _set_SV39_Paddr_PPNi : register_ref regstate register_value SV39_Paddr -> mword ty44 -> M unit\\ \ \\val _set_SV48_PTE_PPNi : register_ref regstate register_value SV48_PTE -> mword ty44 -> M unit\\ \ \\val _set_SV48_Paddr_PPNi : register_ref regstate register_value SV48_Paddr -> mword ty44 -> M unit\\ definition get_SV32_Paddr_PgOfs :: \ SV32_Paddr \(12)Word.word \ where \ get_SV32_Paddr_PgOfs v = ( (subrange_vec_dec(SV32_Paddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) :: 12 Word.word))\ for v :: " SV32_Paddr " definition set_SV32_Paddr_PgOfs :: \((regstate),(register_value),(SV32_Paddr))register_ref \(12)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_Paddr_PgOfs r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_Paddr_bits := ((update_subrange_vec_dec(SV32_Paddr_bits r) (( 11 :: int)::ii) (( 0 :: int)::ii) v :: 34 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_Paddr))register_ref " and v :: "(12)Word.word " definition update_SV32_Paddr_PgOfs :: \ SV32_Paddr \(12)Word.word \ SV32_Paddr \ where \ update_SV32_Paddr_PgOfs v x = ( ( v (| SV32_Paddr_bits := ((update_subrange_vec_dec(SV32_Paddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) x :: 34 Word.word)) |)))\ for v :: " SV32_Paddr " and x :: "(12)Word.word " \ \\val undefined_SV32_PTE : unit -> M SV32_PTE\\ definition undefined_SV32_PTE :: \ unit \((register_value),(SV32_PTE),(exception))monad \ where \ undefined_SV32_PTE _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((| SV32_PTE_bits = w__0 |)))))\ \ \\val Mk_SV32_PTE : mword ty32 -> SV32_PTE\\ definition Mk_SV32_PTE :: \(32)Word.word \ SV32_PTE \ where \ Mk_SV32_PTE v = ( (| SV32_PTE_bits = v |) )\ for v :: "(32)Word.word " definition get_SV32_PTE_bits :: \ SV32_PTE \(32)Word.word \ where \ get_SV32_PTE_bits v = ( (subrange_vec_dec(SV32_PTE_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for v :: " SV32_PTE " definition set_SV32_PTE_bits :: \((regstate),(register_value),(SV32_PTE))register_ref \(32)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_PTE_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_PTE_bits := ((update_subrange_vec_dec(SV32_PTE_bits r) (( 31 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_PTE))register_ref " and v :: "(32)Word.word " definition update_SV32_PTE_bits :: \ SV32_PTE \(32)Word.word \ SV32_PTE \ where \ update_SV32_PTE_bits v x = ( ( v (| SV32_PTE_bits := ((update_subrange_vec_dec(SV32_PTE_bits v) (( 31 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " SV32_PTE " and x :: "(32)Word.word " \ \\val _get_SV32_PTE_BITS : SV32_PTE -> mword ty8\\ definition get_SV32_PTE_BITS :: \ SV32_PTE \(8)Word.word \ where \ get_SV32_PTE_BITS v = ( (subrange_vec_dec(SV32_PTE_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word))\ for v :: " SV32_PTE " \ \\val _set_SV32_PTE_BITS : register_ref regstate register_value SV32_PTE -> mword ty8 -> M unit\\ definition set_SV32_PTE_BITS :: \((regstate),(register_value),(SV32_PTE))register_ref \(8)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_PTE_BITS r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_PTE_bits := ((update_subrange_vec_dec(SV32_PTE_bits r) (( 7 :: int)::ii) (( 0 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_PTE))register_ref " and v :: "(8)Word.word " \ \\val _update_SV32_PTE_BITS : SV32_PTE -> mword ty8 -> SV32_PTE\\ definition update_SV32_PTE_BITS :: \ SV32_PTE \(8)Word.word \ SV32_PTE \ where \ update_SV32_PTE_BITS v x = ( ( v (| SV32_PTE_bits := ((update_subrange_vec_dec(SV32_PTE_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " SV32_PTE " and x :: "(8)Word.word " \ \\val _update_SV39_PTE_BITS : SV39_PTE -> mword ty8 -> SV39_PTE\\ \ \\val _update_SV48_PTE_BITS : SV48_PTE -> mword ty8 -> SV48_PTE\\ \ \\val _get_SV39_PTE_BITS : SV39_PTE -> mword ty8\\ \ \\val _get_SV48_PTE_BITS : SV48_PTE -> mword ty8\\ \ \\val _set_SV39_PTE_BITS : register_ref regstate register_value SV39_PTE -> mword ty8 -> M unit\\ \ \\val _set_SV48_PTE_BITS : register_ref regstate register_value SV48_PTE -> mword ty8 -> M unit\\ definition get_SV32_PTE_PPNi :: \ SV32_PTE \(22)Word.word \ where \ get_SV32_PTE_PPNi v = ( (subrange_vec_dec(SV32_PTE_bits v) (( 31 :: int)::ii) (( 10 :: int)::ii) :: 22 Word.word))\ for v :: " SV32_PTE " definition set_SV32_PTE_PPNi :: \((regstate),(register_value),(SV32_PTE))register_ref \(22)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_PTE_PPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_PTE_bits := ((update_subrange_vec_dec(SV32_PTE_bits r) (( 31 :: int)::ii) (( 10 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_PTE))register_ref " and v :: "(22)Word.word " definition update_SV32_PTE_PPNi :: \ SV32_PTE \(22)Word.word \ SV32_PTE \ where \ update_SV32_PTE_PPNi v x = ( ( v (| SV32_PTE_bits := ((update_subrange_vec_dec(SV32_PTE_bits v) (( 31 :: int)::ii) (( 10 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " SV32_PTE " and x :: "(22)Word.word " \ \\val _get_SV32_PTE_RSW : SV32_PTE -> mword ty2\\ definition get_SV32_PTE_RSW :: \ SV32_PTE \(2)Word.word \ where \ get_SV32_PTE_RSW v = ( (subrange_vec_dec(SV32_PTE_bits v) (( 9 :: int)::ii) (( 8 :: int)::ii) :: 2 Word.word))\ for v :: " SV32_PTE " \ \\val _set_SV32_PTE_RSW : register_ref regstate register_value SV32_PTE -> mword ty2 -> M unit\\ definition set_SV32_PTE_RSW :: \((regstate),(register_value),(SV32_PTE))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV32_PTE_RSW r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV32_PTE_bits := ((update_subrange_vec_dec(SV32_PTE_bits r) (( 9 :: int)::ii) (( 8 :: int)::ii) v :: 32 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV32_PTE))register_ref " and v :: "(2)Word.word " \ \\val _update_SV32_PTE_RSW : SV32_PTE -> mword ty2 -> SV32_PTE\\ definition update_SV32_PTE_RSW :: \ SV32_PTE \(2)Word.word \ SV32_PTE \ where \ update_SV32_PTE_RSW v x = ( ( v (| SV32_PTE_bits := ((update_subrange_vec_dec(SV32_PTE_bits v) (( 9 :: int)::ii) (( 8 :: int)::ii) x :: 32 Word.word)) |)))\ for v :: " SV32_PTE " and x :: "(2)Word.word " \ \\val _update_SV39_PTE_RSW : SV39_PTE -> mword ty2 -> SV39_PTE\\ \ \\val _update_SV48_PTE_RSW : SV48_PTE -> mword ty2 -> SV48_PTE\\ \ \\val _get_SV39_PTE_RSW : SV39_PTE -> mword ty2\\ \ \\val _get_SV48_PTE_RSW : SV48_PTE -> mword ty2\\ \ \\val _set_SV39_PTE_RSW : register_ref regstate register_value SV39_PTE -> mword ty2 -> M unit\\ \ \\val _set_SV48_PTE_RSW : register_ref regstate register_value SV48_PTE -> mword ty2 -> M unit\\ \ \\val curAsid64 : mword ty64 -> mword ty16\\ definition curAsid64 :: \(64)Word.word \(16)Word.word \ where \ curAsid64 satp1 = ( (let s = (Mk_Satp64 satp1) in (get_Satp64_Asid s :: 16 Word.word)))\ for satp1 :: "(64)Word.word " \ \\val curPTB64 : mword ty64 -> mword ty56\\ definition curPTB64 :: \(64)Word.word \(56)Word.word \ where \ curPTB64 satp1 = ( (let s = (Mk_Satp64 satp1) in (shiftl ((EXTZ (( 56 :: int)::ii) ((get_Satp64_PPN s :: 44 Word.word)) :: 56 Word.word)) PAGESIZE_BITS :: 56 Word.word)))\ for satp1 :: "(64)Word.word " definition SV39_LEVEL_BITS :: \ int \ where \ SV39_LEVEL_BITS = ( (( 9 :: int)::ii))\ definition SV39_LEVELS :: \ int \ where \ SV39_LEVELS = ( (( 3 :: int)::ii))\ definition PTE39_LOG_SIZE :: \ int \ where \ PTE39_LOG_SIZE = ( (( 3 :: int)::ii))\ definition PTE39_SIZE :: \ int \ where \ PTE39_SIZE = ( (( 8 :: int)::ii))\ \ \\val undefined_SV39_Vaddr : unit -> M SV39_Vaddr\\ definition undefined_SV39_Vaddr :: \ unit \((register_value),(SV39_Vaddr),(exception))monad \ where \ undefined_SV39_Vaddr _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 39 Word.word) M) \ ((\ (w__0 :: 39 Word.word) . return ((| SV39_Vaddr_bits = w__0 |)))))\ \ \\val Mk_SV39_Vaddr : mword ty39 -> SV39_Vaddr\\ definition Mk_SV39_Vaddr :: \(39)Word.word \ SV39_Vaddr \ where \ Mk_SV39_Vaddr v = ( (| SV39_Vaddr_bits = v |) )\ for v :: "(39)Word.word " definition get_SV39_Vaddr_bits :: \ SV39_Vaddr \(39)Word.word \ where \ get_SV39_Vaddr_bits v = ( (subrange_vec_dec(SV39_Vaddr_bits v) (( 38 :: int)::ii) (( 0 :: int)::ii) :: 39 Word.word))\ for v :: " SV39_Vaddr " definition set_SV39_Vaddr_bits :: \((regstate),(register_value),(SV39_Vaddr))register_ref \(39)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_Vaddr_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_Vaddr_bits := ((update_subrange_vec_dec(SV39_Vaddr_bits r) (( 38 :: int)::ii) (( 0 :: int)::ii) v :: 39 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_Vaddr))register_ref " and v :: "(39)Word.word " definition update_SV39_Vaddr_bits :: \ SV39_Vaddr \(39)Word.word \ SV39_Vaddr \ where \ update_SV39_Vaddr_bits v x = ( ( v (| SV39_Vaddr_bits := ((update_subrange_vec_dec(SV39_Vaddr_bits v) (( 38 :: int)::ii) (( 0 :: int)::ii) x :: 39 Word.word)) |)))\ for v :: " SV39_Vaddr " and x :: "(39)Word.word " definition get_SV39_Vaddr_PgOfs :: \ SV39_Vaddr \(12)Word.word \ where \ get_SV39_Vaddr_PgOfs v = ( (subrange_vec_dec(SV39_Vaddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) :: 12 Word.word))\ for v :: " SV39_Vaddr " definition set_SV39_Vaddr_PgOfs :: \((regstate),(register_value),(SV39_Vaddr))register_ref \(12)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_Vaddr_PgOfs r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_Vaddr_bits := ((update_subrange_vec_dec(SV39_Vaddr_bits r) (( 11 :: int)::ii) (( 0 :: int)::ii) v :: 39 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_Vaddr))register_ref " and v :: "(12)Word.word " definition update_SV39_Vaddr_PgOfs :: \ SV39_Vaddr \(12)Word.word \ SV39_Vaddr \ where \ update_SV39_Vaddr_PgOfs v x = ( ( v (| SV39_Vaddr_bits := ((update_subrange_vec_dec(SV39_Vaddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) x :: 39 Word.word)) |)))\ for v :: " SV39_Vaddr " and x :: "(12)Word.word " definition get_SV39_Vaddr_VPNi :: \ SV39_Vaddr \(27)Word.word \ where \ get_SV39_Vaddr_VPNi v = ( (subrange_vec_dec(SV39_Vaddr_bits v) (( 38 :: int)::ii) (( 12 :: int)::ii) :: 27 Word.word))\ for v :: " SV39_Vaddr " definition set_SV39_Vaddr_VPNi :: \((regstate),(register_value),(SV39_Vaddr))register_ref \(27)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_Vaddr_VPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_Vaddr_bits := ((update_subrange_vec_dec(SV39_Vaddr_bits r) (( 38 :: int)::ii) (( 12 :: int)::ii) v :: 39 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_Vaddr))register_ref " and v :: "(27)Word.word " definition update_SV39_Vaddr_VPNi :: \ SV39_Vaddr \(27)Word.word \ SV39_Vaddr \ where \ update_SV39_Vaddr_VPNi v x = ( ( v (| SV39_Vaddr_bits := ((update_subrange_vec_dec(SV39_Vaddr_bits v) (( 38 :: int)::ii) (( 12 :: int)::ii) x :: 39 Word.word)) |)))\ for v :: " SV39_Vaddr " and x :: "(27)Word.word " \ \\val undefined_SV39_Paddr : unit -> M SV39_Paddr\\ definition undefined_SV39_Paddr :: \ unit \((register_value),(SV39_Paddr),(exception))monad \ where \ undefined_SV39_Paddr _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 56 Word.word) M) \ ((\ (w__0 :: 56 Word.word) . return ((| SV39_Paddr_bits = w__0 |)))))\ \ \\val Mk_SV39_Paddr : mword ty56 -> SV39_Paddr\\ definition Mk_SV39_Paddr :: \(56)Word.word \ SV39_Paddr \ where \ Mk_SV39_Paddr v = ( (| SV39_Paddr_bits = v |) )\ for v :: "(56)Word.word " definition get_SV39_Paddr_bits :: \ SV39_Paddr \(56)Word.word \ where \ get_SV39_Paddr_bits v = ( (subrange_vec_dec(SV39_Paddr_bits v) (( 55 :: int)::ii) (( 0 :: int)::ii) :: 56 Word.word))\ for v :: " SV39_Paddr " definition set_SV39_Paddr_bits :: \((regstate),(register_value),(SV39_Paddr))register_ref \(56)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_Paddr_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_Paddr_bits := ((update_subrange_vec_dec(SV39_Paddr_bits r) (( 55 :: int)::ii) (( 0 :: int)::ii) v :: 56 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_Paddr))register_ref " and v :: "(56)Word.word " definition update_SV39_Paddr_bits :: \ SV39_Paddr \(56)Word.word \ SV39_Paddr \ where \ update_SV39_Paddr_bits v x = ( ( v (| SV39_Paddr_bits := ((update_subrange_vec_dec(SV39_Paddr_bits v) (( 55 :: int)::ii) (( 0 :: int)::ii) x :: 56 Word.word)) |)))\ for v :: " SV39_Paddr " and x :: "(56)Word.word " definition get_SV39_Paddr_PPNi :: \ SV39_Paddr \(44)Word.word \ where \ get_SV39_Paddr_PPNi v = ( (subrange_vec_dec(SV39_Paddr_bits v) (( 55 :: int)::ii) (( 12 :: int)::ii) :: 44 Word.word))\ for v :: " SV39_Paddr " definition set_SV39_Paddr_PPNi :: \((regstate),(register_value),(SV39_Paddr))register_ref \(44)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_Paddr_PPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_Paddr_bits := ((update_subrange_vec_dec(SV39_Paddr_bits r) (( 55 :: int)::ii) (( 12 :: int)::ii) v :: 56 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_Paddr))register_ref " and v :: "(44)Word.word " definition update_SV39_Paddr_PPNi :: \ SV39_Paddr \(44)Word.word \ SV39_Paddr \ where \ update_SV39_Paddr_PPNi v x = ( ( v (| SV39_Paddr_bits := ((update_subrange_vec_dec(SV39_Paddr_bits v) (( 55 :: int)::ii) (( 12 :: int)::ii) x :: 56 Word.word)) |)))\ for v :: " SV39_Paddr " and x :: "(44)Word.word " definition get_SV39_Paddr_PgOfs :: \ SV39_Paddr \(12)Word.word \ where \ get_SV39_Paddr_PgOfs v = ( (subrange_vec_dec(SV39_Paddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) :: 12 Word.word))\ for v :: " SV39_Paddr " definition set_SV39_Paddr_PgOfs :: \((regstate),(register_value),(SV39_Paddr))register_ref \(12)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_Paddr_PgOfs r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_Paddr_bits := ((update_subrange_vec_dec(SV39_Paddr_bits r) (( 11 :: int)::ii) (( 0 :: int)::ii) v :: 56 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_Paddr))register_ref " and v :: "(12)Word.word " definition update_SV39_Paddr_PgOfs :: \ SV39_Paddr \(12)Word.word \ SV39_Paddr \ where \ update_SV39_Paddr_PgOfs v x = ( ( v (| SV39_Paddr_bits := ((update_subrange_vec_dec(SV39_Paddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) x :: 56 Word.word)) |)))\ for v :: " SV39_Paddr " and x :: "(12)Word.word " \ \\val undefined_SV39_PTE : unit -> M SV39_PTE\\ definition undefined_SV39_PTE :: \ unit \((register_value),(SV39_PTE),(exception))monad \ where \ undefined_SV39_PTE _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . return ((| SV39_PTE_bits = w__0 |)))))\ \ \\val Mk_SV39_PTE : mword ty64 -> SV39_PTE\\ definition Mk_SV39_PTE :: \(64)Word.word \ SV39_PTE \ where \ Mk_SV39_PTE v = ( (| SV39_PTE_bits = v |) )\ for v :: "(64)Word.word " definition get_SV39_PTE_bits :: \ SV39_PTE \(64)Word.word \ where \ get_SV39_PTE_bits v = ( (subrange_vec_dec(SV39_PTE_bits v) (( 63 :: int)::ii) (( 0 :: int)::ii) :: 64 Word.word))\ for v :: " SV39_PTE " definition set_SV39_PTE_bits :: \((regstate),(register_value),(SV39_PTE))register_ref \(64)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_PTE_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits r) (( 63 :: int)::ii) (( 0 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_PTE))register_ref " and v :: "(64)Word.word " definition update_SV39_PTE_bits :: \ SV39_PTE \(64)Word.word \ SV39_PTE \ where \ update_SV39_PTE_bits v x = ( ( v (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits v) (( 63 :: int)::ii) (( 0 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV39_PTE " and x :: "(64)Word.word " definition get_SV39_PTE_BITS :: \ SV39_PTE \(8)Word.word \ where \ get_SV39_PTE_BITS v = ( (subrange_vec_dec(SV39_PTE_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word))\ for v :: " SV39_PTE " definition set_SV39_PTE_BITS :: \((regstate),(register_value),(SV39_PTE))register_ref \(8)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_PTE_BITS r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits r) (( 7 :: int)::ii) (( 0 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_PTE))register_ref " and v :: "(8)Word.word " definition update_SV39_PTE_BITS :: \ SV39_PTE \(8)Word.word \ SV39_PTE \ where \ update_SV39_PTE_BITS v x = ( ( v (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV39_PTE " and x :: "(8)Word.word " \ \\val _get_SV39_PTE_Ext : SV39_PTE -> mword ty10\\ definition get_SV39_PTE_Ext :: \ SV39_PTE \(10)Word.word \ where \ get_SV39_PTE_Ext v = ( (subrange_vec_dec(SV39_PTE_bits v) (( 63 :: int)::ii) (( 54 :: int)::ii) :: 10 Word.word))\ for v :: " SV39_PTE " \ \\val _set_SV39_PTE_Ext : register_ref regstate register_value SV39_PTE -> mword ty10 -> M unit\\ definition set_SV39_PTE_Ext :: \((regstate),(register_value),(SV39_PTE))register_ref \(10)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_PTE_Ext r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits r) (( 63 :: int)::ii) (( 54 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_PTE))register_ref " and v :: "(10)Word.word " \ \\val _update_SV39_PTE_Ext : SV39_PTE -> mword ty10 -> SV39_PTE\\ definition update_SV39_PTE_Ext :: \ SV39_PTE \(10)Word.word \ SV39_PTE \ where \ update_SV39_PTE_Ext v x = ( ( v (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits v) (( 63 :: int)::ii) (( 54 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV39_PTE " and x :: "(10)Word.word " \ \\val _update_SV48_PTE_Ext : SV48_PTE -> mword ty10 -> SV48_PTE\\ \ \\val _get_SV48_PTE_Ext : SV48_PTE -> mword ty10\\ \ \\val _set_SV48_PTE_Ext : register_ref regstate register_value SV48_PTE -> mword ty10 -> M unit\\ definition get_SV39_PTE_PPNi :: \ SV39_PTE \(44)Word.word \ where \ get_SV39_PTE_PPNi v = ( (subrange_vec_dec(SV39_PTE_bits v) (( 53 :: int)::ii) (( 10 :: int)::ii) :: 44 Word.word))\ for v :: " SV39_PTE " definition set_SV39_PTE_PPNi :: \((regstate),(register_value),(SV39_PTE))register_ref \(44)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_PTE_PPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits r) (( 53 :: int)::ii) (( 10 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_PTE))register_ref " and v :: "(44)Word.word " definition update_SV39_PTE_PPNi :: \ SV39_PTE \(44)Word.word \ SV39_PTE \ where \ update_SV39_PTE_PPNi v x = ( ( v (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits v) (( 53 :: int)::ii) (( 10 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV39_PTE " and x :: "(44)Word.word " definition get_SV39_PTE_RSW :: \ SV39_PTE \(2)Word.word \ where \ get_SV39_PTE_RSW v = ( (subrange_vec_dec(SV39_PTE_bits v) (( 9 :: int)::ii) (( 8 :: int)::ii) :: 2 Word.word))\ for v :: " SV39_PTE " definition set_SV39_PTE_RSW :: \((regstate),(register_value),(SV39_PTE))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV39_PTE_RSW r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits r) (( 9 :: int)::ii) (( 8 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV39_PTE))register_ref " and v :: "(2)Word.word " definition update_SV39_PTE_RSW :: \ SV39_PTE \(2)Word.word \ SV39_PTE \ where \ update_SV39_PTE_RSW v x = ( ( v (| SV39_PTE_bits := ((update_subrange_vec_dec(SV39_PTE_bits v) (( 9 :: int)::ii) (( 8 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV39_PTE " and x :: "(2)Word.word " definition SV48_LEVEL_BITS :: \ int \ where \ SV48_LEVEL_BITS = ( (( 9 :: int)::ii))\ definition SV48_LEVELS :: \ int \ where \ SV48_LEVELS = ( (( 4 :: int)::ii))\ definition PTE48_LOG_SIZE :: \ int \ where \ PTE48_LOG_SIZE = ( (( 3 :: int)::ii))\ definition PTE48_SIZE :: \ int \ where \ PTE48_SIZE = ( (( 8 :: int)::ii))\ \ \\val undefined_SV48_Vaddr : unit -> M SV48_Vaddr\\ definition undefined_SV48_Vaddr :: \ unit \((register_value),(SV48_Vaddr),(exception))monad \ where \ undefined_SV48_Vaddr _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 48 Word.word) M) \ ((\ (w__0 :: 48 Word.word) . return ((| SV48_Vaddr_bits = w__0 |)))))\ \ \\val Mk_SV48_Vaddr : mword ty48 -> SV48_Vaddr\\ definition Mk_SV48_Vaddr :: \(48)Word.word \ SV48_Vaddr \ where \ Mk_SV48_Vaddr v = ( (| SV48_Vaddr_bits = v |) )\ for v :: "(48)Word.word " definition get_SV48_Vaddr_bits :: \ SV48_Vaddr \(48)Word.word \ where \ get_SV48_Vaddr_bits v = ( (subrange_vec_dec(SV48_Vaddr_bits v) (( 47 :: int)::ii) (( 0 :: int)::ii) :: 48 Word.word))\ for v :: " SV48_Vaddr " definition set_SV48_Vaddr_bits :: \((regstate),(register_value),(SV48_Vaddr))register_ref \(48)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_Vaddr_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_Vaddr_bits := ((update_subrange_vec_dec(SV48_Vaddr_bits r) (( 47 :: int)::ii) (( 0 :: int)::ii) v :: 48 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_Vaddr))register_ref " and v :: "(48)Word.word " definition update_SV48_Vaddr_bits :: \ SV48_Vaddr \(48)Word.word \ SV48_Vaddr \ where \ update_SV48_Vaddr_bits v x = ( ( v (| SV48_Vaddr_bits := ((update_subrange_vec_dec(SV48_Vaddr_bits v) (( 47 :: int)::ii) (( 0 :: int)::ii) x :: 48 Word.word)) |)))\ for v :: " SV48_Vaddr " and x :: "(48)Word.word " definition get_SV48_Vaddr_PgOfs :: \ SV48_Vaddr \(12)Word.word \ where \ get_SV48_Vaddr_PgOfs v = ( (subrange_vec_dec(SV48_Vaddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) :: 12 Word.word))\ for v :: " SV48_Vaddr " definition set_SV48_Vaddr_PgOfs :: \((regstate),(register_value),(SV48_Vaddr))register_ref \(12)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_Vaddr_PgOfs r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_Vaddr_bits := ((update_subrange_vec_dec(SV48_Vaddr_bits r) (( 11 :: int)::ii) (( 0 :: int)::ii) v :: 48 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_Vaddr))register_ref " and v :: "(12)Word.word " definition update_SV48_Vaddr_PgOfs :: \ SV48_Vaddr \(12)Word.word \ SV48_Vaddr \ where \ update_SV48_Vaddr_PgOfs v x = ( ( v (| SV48_Vaddr_bits := ((update_subrange_vec_dec(SV48_Vaddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) x :: 48 Word.word)) |)))\ for v :: " SV48_Vaddr " and x :: "(12)Word.word " definition get_SV48_Vaddr_VPNi :: \ SV48_Vaddr \(27)Word.word \ where \ get_SV48_Vaddr_VPNi v = ( (subrange_vec_dec(SV48_Vaddr_bits v) (( 38 :: int)::ii) (( 12 :: int)::ii) :: 27 Word.word))\ for v :: " SV48_Vaddr " definition set_SV48_Vaddr_VPNi :: \((regstate),(register_value),(SV48_Vaddr))register_ref \(27)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_Vaddr_VPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_Vaddr_bits := ((update_subrange_vec_dec(SV48_Vaddr_bits r) (( 38 :: int)::ii) (( 12 :: int)::ii) v :: 48 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_Vaddr))register_ref " and v :: "(27)Word.word " definition update_SV48_Vaddr_VPNi :: \ SV48_Vaddr \(27)Word.word \ SV48_Vaddr \ where \ update_SV48_Vaddr_VPNi v x = ( ( v (| SV48_Vaddr_bits := ((update_subrange_vec_dec(SV48_Vaddr_bits v) (( 38 :: int)::ii) (( 12 :: int)::ii) x :: 48 Word.word)) |)))\ for v :: " SV48_Vaddr " and x :: "(27)Word.word " \ \\val undefined_SV48_Paddr : unit -> M SV48_Paddr\\ definition undefined_SV48_Paddr :: \ unit \((register_value),(SV48_Paddr),(exception))monad \ where \ undefined_SV48_Paddr _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 56 Word.word) M) \ ((\ (w__0 :: 56 Word.word) . return ((| SV48_Paddr_bits = w__0 |)))))\ \ \\val Mk_SV48_Paddr : mword ty56 -> SV48_Paddr\\ definition Mk_SV48_Paddr :: \(56)Word.word \ SV48_Paddr \ where \ Mk_SV48_Paddr v = ( (| SV48_Paddr_bits = v |) )\ for v :: "(56)Word.word " definition get_SV48_Paddr_bits :: \ SV48_Paddr \(56)Word.word \ where \ get_SV48_Paddr_bits v = ( (subrange_vec_dec(SV48_Paddr_bits v) (( 55 :: int)::ii) (( 0 :: int)::ii) :: 56 Word.word))\ for v :: " SV48_Paddr " definition set_SV48_Paddr_bits :: \((regstate),(register_value),(SV48_Paddr))register_ref \(56)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_Paddr_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_Paddr_bits := ((update_subrange_vec_dec(SV48_Paddr_bits r) (( 55 :: int)::ii) (( 0 :: int)::ii) v :: 56 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_Paddr))register_ref " and v :: "(56)Word.word " definition update_SV48_Paddr_bits :: \ SV48_Paddr \(56)Word.word \ SV48_Paddr \ where \ update_SV48_Paddr_bits v x = ( ( v (| SV48_Paddr_bits := ((update_subrange_vec_dec(SV48_Paddr_bits v) (( 55 :: int)::ii) (( 0 :: int)::ii) x :: 56 Word.word)) |)))\ for v :: " SV48_Paddr " and x :: "(56)Word.word " definition get_SV48_Paddr_PPNi :: \ SV48_Paddr \(44)Word.word \ where \ get_SV48_Paddr_PPNi v = ( (subrange_vec_dec(SV48_Paddr_bits v) (( 55 :: int)::ii) (( 12 :: int)::ii) :: 44 Word.word))\ for v :: " SV48_Paddr " definition set_SV48_Paddr_PPNi :: \((regstate),(register_value),(SV48_Paddr))register_ref \(44)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_Paddr_PPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_Paddr_bits := ((update_subrange_vec_dec(SV48_Paddr_bits r) (( 55 :: int)::ii) (( 12 :: int)::ii) v :: 56 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_Paddr))register_ref " and v :: "(44)Word.word " definition update_SV48_Paddr_PPNi :: \ SV48_Paddr \(44)Word.word \ SV48_Paddr \ where \ update_SV48_Paddr_PPNi v x = ( ( v (| SV48_Paddr_bits := ((update_subrange_vec_dec(SV48_Paddr_bits v) (( 55 :: int)::ii) (( 12 :: int)::ii) x :: 56 Word.word)) |)))\ for v :: " SV48_Paddr " and x :: "(44)Word.word " definition get_SV48_Paddr_PgOfs :: \ SV48_Paddr \(12)Word.word \ where \ get_SV48_Paddr_PgOfs v = ( (subrange_vec_dec(SV48_Paddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) :: 12 Word.word))\ for v :: " SV48_Paddr " definition set_SV48_Paddr_PgOfs :: \((regstate),(register_value),(SV48_Paddr))register_ref \(12)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_Paddr_PgOfs r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_Paddr_bits := ((update_subrange_vec_dec(SV48_Paddr_bits r) (( 11 :: int)::ii) (( 0 :: int)::ii) v :: 56 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_Paddr))register_ref " and v :: "(12)Word.word " definition update_SV48_Paddr_PgOfs :: \ SV48_Paddr \(12)Word.word \ SV48_Paddr \ where \ update_SV48_Paddr_PgOfs v x = ( ( v (| SV48_Paddr_bits := ((update_subrange_vec_dec(SV48_Paddr_bits v) (( 11 :: int)::ii) (( 0 :: int)::ii) x :: 56 Word.word)) |)))\ for v :: " SV48_Paddr " and x :: "(12)Word.word " \ \\val undefined_SV48_PTE : unit -> M SV48_PTE\\ definition undefined_SV48_PTE :: \ unit \((register_value),(SV48_PTE),(exception))monad \ where \ undefined_SV48_PTE _ = ( ((return (failwith (''undefined value of unsupported type''))) :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . return ((| SV48_PTE_bits = w__0 |)))))\ \ \\val Mk_SV48_PTE : mword ty64 -> SV48_PTE\\ definition Mk_SV48_PTE :: \(64)Word.word \ SV48_PTE \ where \ Mk_SV48_PTE v = ( (| SV48_PTE_bits = v |) )\ for v :: "(64)Word.word " definition get_SV48_PTE_bits :: \ SV48_PTE \(64)Word.word \ where \ get_SV48_PTE_bits v = ( (subrange_vec_dec(SV48_PTE_bits v) (( 63 :: int)::ii) (( 0 :: int)::ii) :: 64 Word.word))\ for v :: " SV48_PTE " definition set_SV48_PTE_bits :: \((regstate),(register_value),(SV48_PTE))register_ref \(64)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_PTE_bits r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits r) (( 63 :: int)::ii) (( 0 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_PTE))register_ref " and v :: "(64)Word.word " definition update_SV48_PTE_bits :: \ SV48_PTE \(64)Word.word \ SV48_PTE \ where \ update_SV48_PTE_bits v x = ( ( v (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits v) (( 63 :: int)::ii) (( 0 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV48_PTE " and x :: "(64)Word.word " definition get_SV48_PTE_BITS :: \ SV48_PTE \(8)Word.word \ where \ get_SV48_PTE_BITS v = ( (subrange_vec_dec(SV48_PTE_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word))\ for v :: " SV48_PTE " definition set_SV48_PTE_BITS :: \((regstate),(register_value),(SV48_PTE))register_ref \(8)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_PTE_BITS r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits r) (( 7 :: int)::ii) (( 0 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_PTE))register_ref " and v :: "(8)Word.word " definition update_SV48_PTE_BITS :: \ SV48_PTE \(8)Word.word \ SV48_PTE \ where \ update_SV48_PTE_BITS v x = ( ( v (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits v) (( 7 :: int)::ii) (( 0 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV48_PTE " and x :: "(8)Word.word " definition get_SV48_PTE_Ext :: \ SV48_PTE \(10)Word.word \ where \ get_SV48_PTE_Ext v = ( (subrange_vec_dec(SV48_PTE_bits v) (( 63 :: int)::ii) (( 54 :: int)::ii) :: 10 Word.word))\ for v :: " SV48_PTE " definition set_SV48_PTE_Ext :: \((regstate),(register_value),(SV48_PTE))register_ref \(10)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_PTE_Ext r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits r) (( 63 :: int)::ii) (( 54 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_PTE))register_ref " and v :: "(10)Word.word " definition update_SV48_PTE_Ext :: \ SV48_PTE \(10)Word.word \ SV48_PTE \ where \ update_SV48_PTE_Ext v x = ( ( v (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits v) (( 63 :: int)::ii) (( 54 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV48_PTE " and x :: "(10)Word.word " definition get_SV48_PTE_PPNi :: \ SV48_PTE \(44)Word.word \ where \ get_SV48_PTE_PPNi v = ( (subrange_vec_dec(SV48_PTE_bits v) (( 53 :: int)::ii) (( 10 :: int)::ii) :: 44 Word.word))\ for v :: " SV48_PTE " definition set_SV48_PTE_PPNi :: \((regstate),(register_value),(SV48_PTE))register_ref \(44)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_PTE_PPNi r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits r) (( 53 :: int)::ii) (( 10 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_PTE))register_ref " and v :: "(44)Word.word " definition update_SV48_PTE_PPNi :: \ SV48_PTE \(44)Word.word \ SV48_PTE \ where \ update_SV48_PTE_PPNi v x = ( ( v (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits v) (( 53 :: int)::ii) (( 10 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV48_PTE " and x :: "(44)Word.word " definition get_SV48_PTE_RSW :: \ SV48_PTE \(2)Word.word \ where \ get_SV48_PTE_RSW v = ( (subrange_vec_dec(SV48_PTE_bits v) (( 9 :: int)::ii) (( 8 :: int)::ii) :: 2 Word.word))\ for v :: " SV48_PTE " definition set_SV48_PTE_RSW :: \((regstate),(register_value),(SV48_PTE))register_ref \(2)Word.word \((register_value),(unit),(exception))monad \ where \ set_SV48_PTE_RSW r_ref v = ( read_reg r_ref \ ((\ r . (let r = (( r (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits r) (( 9 :: int)::ii) (( 8 :: int)::ii) v :: 64 Word.word)) |))) in write_reg r_ref r))))\ for r_ref :: "((regstate),(register_value),(SV48_PTE))register_ref " and v :: "(2)Word.word " definition update_SV48_PTE_RSW :: \ SV48_PTE \(2)Word.word \ SV48_PTE \ where \ update_SV48_PTE_RSW v x = ( ( v (| SV48_PTE_bits := ((update_subrange_vec_dec(SV48_PTE_bits v) (( 9 :: int)::ii) (( 8 :: int)::ii) x :: 64 Word.word)) |)))\ for v :: " SV48_PTE " and x :: "(2)Word.word " \ \\val make_TLB_Entry : forall 'asidlen 'palen 'ptelen 'valen. Size 'asidlen, Size 'palen, Size 'ptelen, Size 'valen => mword 'asidlen -> bool -> mword 'valen -> mword 'palen -> mword 'ptelen -> ii -> mword 'palen -> ii -> M (TLB_Entry 'asidlen 'valen 'palen 'ptelen)\\ definition make_TLB_Entry :: \('asidlen::len)Word.word \ bool \('valen::len)Word.word \('palen::len)Word.word \('ptelen::len)Word.word \ int \('palen::len)Word.word \ int \((register_value),((('asidlen::len),('valen::len),('palen::len),('ptelen::len))TLB_Entry),(exception))monad \ where \ make_TLB_Entry asid global1 vAddr pAddr pte level pteAddr levelBitSize = ( (let (shift :: ii) = (PAGESIZE_BITS + ((level * levelBitSize))) in (let (vAddrMask :: ( 'valen::len)Word.word) = ((sub_vec_int ((shiftl ((xor_vec vAddr ((xor_vec vAddr ((EXTZ ((int (size vAddr))) ( 0b1 :: 1 Word.word) :: ( 'valen::len)Word.word)) :: ( 'valen::len)Word.word)) :: ( 'valen::len)Word.word)) shift :: ( 'valen::len)Word.word)) (( 1 :: int)::ii) :: ( 'valen::len)Word.word)) in (let (vMatchMask :: ( 'valen::len)Word.word) = ((not_vec vAddrMask :: ( 'valen::len)Word.word)) in (read_reg mcycle_ref :: ( 64 Word.word) M) \ ((\ (w__0 :: 64 Word.word) . return ((| TLB_Entry_asid = asid, TLB_Entry_global = global1, TLB_Entry_vAddr = ((and_vec vAddr vMatchMask :: ( 'valen::len)Word.word)), TLB_Entry_pAddr = ((shiftl ((shiftr pAddr shift :: ( 'palen::len)Word.word)) shift :: ( 'palen::len)Word.word)), TLB_Entry_vMatchMask = vMatchMask, TLB_Entry_vAddrMask = vAddrMask, TLB_Entry_pte = pte, TLB_Entry_pteAddr = pteAddr, TLB_Entry_age = w__0 |))))))))\ for asid :: "('asidlen::len)Word.word " and global1 :: " bool " and vAddr :: "('valen::len)Word.word " and pAddr :: "('palen::len)Word.word " and pte :: "('ptelen::len)Word.word " and level :: " int " and pteAddr :: "('palen::len)Word.word " and levelBitSize :: " int " \ \\val match_TLB_Entry : forall 'asidlen 'palen 'ptelen 'valen. Size 'asidlen, Size 'valen => TLB_Entry 'asidlen 'valen 'palen 'ptelen -> mword 'asidlen -> mword 'valen -> bool\\ definition match_TLB_Entry :: \(('asidlen::len),('valen::len),('palen::len),('ptelen::len))TLB_Entry \('asidlen::len)Word.word \('valen::len)Word.word \ bool \ where \ match_TLB_Entry ent asid vaddr = ( (((((TLB_Entry_global ent) \ ((((TLB_Entry_asid ent) = asid)))))) \ ((((TLB_Entry_vAddr ent) = ((and_vec(TLB_Entry_vMatchMask ent) vaddr :: ( 'valen::len)Word.word)))))))\ for ent :: "(('asidlen::len),('valen::len),('palen::len),('ptelen::len))TLB_Entry " and asid :: "('asidlen::len)Word.word " and vaddr :: "('valen::len)Word.word " \ \\val flush_TLB_Entry : forall 'asidlen 'palen 'ptelen 'valen. Size 'asidlen, Size 'valen => TLB_Entry 'asidlen 'valen 'palen 'ptelen -> maybe (mword 'asidlen) -> maybe (mword 'valen) -> bool\\ fun flush_TLB_Entry :: \(('asidlen::len),('valen::len),('palen::len),('ptelen::len))TLB_Entry \(('asidlen::len)Word.word)option \(('valen::len)Word.word)option \ bool \ where \ flush_TLB_Entry e None None = ( True )\ for e :: "(('asidlen::len),('valen::len),('palen::len),('ptelen::len))TLB_Entry " |\ flush_TLB_Entry e None (Some (a)) = ( ((TLB_Entry_vAddr e) = ((and_vec(TLB_Entry_vMatchMask e) a :: ( 'valen::len)Word.word))))\ for e :: "(('asidlen::len),('valen::len),('palen::len),('ptelen::len))TLB_Entry " and a :: "('valen::len)Word.word " |\ flush_TLB_Entry e (Some (i)) None = ( (((((TLB_Entry_asid e) = i))) \ ((\(TLB_Entry_global e)))))\ for e :: "(('asidlen::len),('valen::len),('palen::len),('ptelen::len))TLB_Entry " and i :: "('asidlen::len)Word.word " |\ flush_TLB_Entry e (Some (i)) (Some (a)) = ( (((((TLB_Entry_asid e) = i))) \ (((((((TLB_Entry_vAddr e) = ((and_vec a(TLB_Entry_vMatchMask e) :: ( 'valen::len)Word.word))))) \ ((\(TLB_Entry_global e))))))))\ for e :: "(('asidlen::len),('valen::len),('palen::len),('ptelen::len))TLB_Entry " and i :: "('asidlen::len)Word.word " and a :: "('valen::len)Word.word " \ \\val to_phys_addr : mword ty34 -> mword ty32\\ definition to_phys_addr :: \(34)Word.word \(32)Word.word \ where \ to_phys_addr a = ( (subrange_vec_dec a (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))\ for a :: "(34)Word.word " \ \\val walk32 : mword ty32 -> AccessType unit -> Privilege -> bool -> bool -> mword ty34 -> ii -> bool -> unit -> M (PTW_Result (mword ty34) SV32_PTE)\\ function (sequential,domintros) walk32 :: \(32)Word.word \(unit)AccessType \ Privilege \ bool \ bool \(34)Word.word \ int \ bool \ unit \((register_value),((((34)Word.word),(SV32_PTE))PTW_Result),(exception))monad \ where \ walk32 vaddr ac priv mxr do_sum ptb level global1 ext_ptw = ( (let va = (Mk_SV32_Vaddr vaddr) in (let (pt_ofs :: paddr32) = ((shiftl ((EXTZ (( 34 :: int)::ii) ((subrange_vec_dec ((shiftr ((get_SV32_Vaddr_VPNi va :: 20 Word.word)) ((level * SV32_LEVEL_BITS)) :: 20 Word.word)) ((SV32_LEVEL_BITS - (( 1 :: int)::ii))) (( 0 :: int)::ii) :: 10 Word.word)) :: 34 Word.word)) PTE32_LOG_SIZE :: 34 Word.word)) in (let pte_addr = ((add_vec ptb pt_ofs :: 34 Word.word)) in (mem_read_priv (Read Data) Supervisor ((to_phys_addr pte_addr :: 32 Word.word)) (( 4 :: int)::ii) False False False :: ( ( 32 Word.word)MemoryOpResult) M) \ ((\ (w__0 :: ( 32 Word.word) MemoryOpResult) . (case w__0 of MemException (_) => return (PTW_Failure (PTW_Access () , ext_ptw)) | MemValue (v) => (let pte = (Mk_SV32_PTE v) in (let pbits = ((get_SV32_PTE_BITS pte :: 8 Word.word)) in (let (ext_pte :: extPte) = default_sv32_ext_pte in (let pattr = (Mk_PTE_Bits pbits) in (let is_global = (global1 \ (((((get_PTE_Bits_G pattr :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))) in if ((isInvalidPTE pbits ext_pte)) then return (PTW_Failure (PTW_Invalid_PTE () , ext_ptw)) else if ((isPTEPtr pbits ext_pte)) then if ((level > (( 0 :: int)::ii))) then (walk32 vaddr ac priv mxr do_sum ((shiftl ((EXTZ (( 34 :: int)::ii) ((get_SV32_PTE_PPNi pte :: 22 Word.word)) :: 34 Word.word)) PAGESIZE_BITS :: 34 Word.word)) ((level - (( 1 :: int)::ii))) is_global ext_ptw :: ( (( 34 Word.word), SV32_PTE)PTW_Result) M) else return (PTW_Failure (PTW_Invalid_PTE () , ext_ptw)) else checkPTEPermission ac priv mxr do_sum pattr ext_pte ext_ptw \ ((\ (w__3 :: PTE_Check) . return ((case w__3 of PTE_Check_Failure ((ext_ptw, ext_ptw_fail)) => PTW_Failure (ext_get_ptw_error ext_ptw_fail, ext_ptw) | PTE_Check_Success (ext_ptw) => if ((level > (( 0 :: int)::ii))) then (let mask1 = ((sub_vec_int ((shiftl ((xor_vec ((get_SV32_PTE_PPNi pte :: 22 Word.word)) ((xor_vec ((get_SV32_PTE_PPNi pte :: 22 Word.word)) ((EXTZ (( 22 :: int)::ii) ( 0b1 :: 1 Word.word) :: 22 Word.word)) :: 22 Word.word)) :: 22 Word.word)) ((level * SV32_LEVEL_BITS)) :: 22 Word.word)) (( 1 :: int)::ii) :: 22 Word.word)) in if (((((and_vec ((get_SV32_PTE_PPNi pte :: 22 Word.word)) mask1 :: 22 Word.word)) \ ((EXTZ (( 22 :: int)::ii) ( 0b0 :: 1 Word.word) :: 22 Word.word))))) then PTW_Failure (PTW_Misaligned () , ext_ptw) else (let ppn = ((or_vec ((get_SV32_PTE_PPNi pte :: 22 Word.word)) ((and_vec ((EXTZ (( 22 :: int)::ii) ((get_SV32_Vaddr_VPNi va :: 20 Word.word)) :: 22 Word.word)) mask1 :: 22 Word.word)) :: 22 Word.word)) in PTW_Success ((concat_vec ppn ((get_SV32_Vaddr_PgOfs va :: 12 Word.word)) :: 34 Word.word), pte, pte_addr, level, is_global, ext_ptw))) else PTW_Success ((concat_vec ((get_SV32_PTE_PPNi pte :: 22 Word.word)) ((get_SV32_Vaddr_PgOfs va :: 12 Word.word)) :: 34 Word.word), pte, pte_addr, level, is_global, ext_ptw) ))))))))) )))))))\ for vaddr :: "(32)Word.word " and ac :: "(unit)AccessType " and priv :: " Privilege " and mxr :: " bool " and do_sum :: " bool " and ptb :: "(34)Word.word " and level :: " int " and global1 :: " bool " and ext_ptw :: " unit " by pat_completeness (auto intro!: let_cong bind_cong MemoryOpResult.case_cong) \ \\val lookup_TLB32 : mword ty9 -> mword ty32 -> M (maybe ((ii * TLB_Entry ty9 ty32 ty34 ty32)))\\ definition lookup_TLB32 :: \(9)Word.word \(32)Word.word \((register_value),((int*((9),(32),(34),(32))TLB_Entry)option),(exception))monad \ where \ lookup_TLB32 asid vaddr = ( read_reg tlb32_ref \ ((\ (w__0 :: ( (9, 32, 34, 32)TLB_Entry)option) . return ((case w__0 of None => None | Some (e) => if ((match_TLB_Entry e asid vaddr)) then Some ((( 0 :: int)::ii), e) else None )))))\ for asid :: "(9)Word.word " and vaddr :: "(32)Word.word " \ \\val add_to_TLB32 : mword ty9 -> mword ty32 -> mword ty34 -> SV32_PTE -> mword ty34 -> ii -> bool -> M unit\\ definition add_to_TLB32 :: \(9)Word.word \(32)Word.word \(34)Word.word \ SV32_PTE \(34)Word.word \ int \ bool \((register_value),(unit),(exception))monad \ where \ add_to_TLB32 asid vAddr pAddr pte pteAddr level global1 = ( make_TLB_Entry asid global1 vAddr pAddr ((get_SV32_PTE_bits pte :: 32 Word.word)) level pteAddr SV32_LEVEL_BITS \ ((\ (ent :: TLB32_Entry) . write_reg tlb32_ref (Some ent))))\ for asid :: "(9)Word.word " and vAddr :: "(32)Word.word " and pAddr :: "(34)Word.word " and pte :: " SV32_PTE " and pteAddr :: "(34)Word.word " and level :: " int " and global1 :: " bool " \ \\val write_TLB32 : ii -> TLB_Entry ty9 ty32 ty34 ty32 -> M unit\\ definition write_TLB32 :: \ int \((9),(32),(34),(32))TLB_Entry \((register_value),(unit),(exception))monad \ where \ write_TLB32 (idx :: ii) (ent :: TLB32_Entry) = ( write_reg tlb32_ref (Some ent))\ for idx :: " int " and ent :: "((9),(32),(34),(32))TLB_Entry " \ \\val flush_TLB32 : maybe (mword ty9) -> maybe (mword ty32) -> M unit\\ definition flush_TLB32 :: \((9)Word.word)option \((32)Word.word)option \((register_value),(unit),(exception))monad \ where \ flush_TLB32 asid addr = ( read_reg tlb32_ref \ ((\ (w__0 :: ( (9, 32, 34, 32)TLB_Entry)option) . (case w__0 of None => return () | Some (e) => if ((flush_TLB_Entry e asid addr)) then write_reg tlb32_ref None else return () ))))\ for asid :: "((9)Word.word)option " and addr :: "((32)Word.word)option " \ \\val translate32 : mword ty9 -> mword ty34 -> mword ty32 -> AccessType unit -> Privilege -> bool -> bool -> ii -> unit -> M (TR_Result (mword ty34) PTW_Error)\\ definition translate32 :: \(9)Word.word \(34)Word.word \(32)Word.word \(unit)AccessType \ Privilege \ bool \ bool \ int \ unit \((register_value),((((34)Word.word),(PTW_Error))TR_Result),(exception))monad \ where \ translate32 asid ptb vAddr ac priv mxr do_sum level ext_ptw = ( lookup_TLB32 asid vAddr \ ((\ (w__0 :: ((ii * (9, 32, 34, 32) TLB_Entry))option) . (case w__0 of Some ((idx, ent)) => (let pte = (Mk_SV32_PTE(TLB_Entry_pte ent)) in (let (ext_pte :: extPte) = ((zeros_implicit (( 10 :: int)::ii) :: 10 Word.word)) in (let pteBits = (Mk_PTE_Bits ((get_SV32_PTE_BITS pte :: 8 Word.word))) in checkPTEPermission ac priv mxr do_sum pteBits ext_pte ext_ptw \ ((\ (w__1 :: PTE_Check) . (case w__1 of PTE_Check_Failure ((ext_ptw, ext_ptw_fail)) => return (TR_Failure (ext_get_ptw_error ext_ptw_fail, ext_ptw)) | PTE_Check_Success (ext_ptw) => (case ((update_PTE_Bits pteBits ac ext_pte :: ((PTE_Bits * 10 Word.word))option)) of None => return (TR_Address ((or_vec(TLB_Entry_pAddr ent) ((EXTZ (( 34 :: int)::ii) ((and_vec vAddr(TLB_Entry_vAddrMask ent) :: 32 Word.word)) :: 34 Word.word)) :: 34 Word.word), ext_ptw)) | Some ((pbits, ext)) => if ((\ ((plat_enable_dirty_update () )))) then return (TR_Failure (PTW_PTE_Update () , ext_ptw)) else (let n_pte = (update_SV32_PTE_BITS pte ((get_PTE_Bits_bits pbits :: 8 Word.word))) in (let (n_ent :: TLB32_Entry) = ent in (let n_ent = (( n_ent (| TLB_Entry_pte := ((get_SV32_PTE_bits n_pte :: 32 Word.word)) |))) in (write_TLB32 idx n_ent \ mem_write_value_priv ((to_phys_addr ((EXTZ (( 34 :: int)::ii)(TLB_Entry_pteAddr ent) :: 34 Word.word)) :: 32 Word.word)) (( 4 :: int)::ii) ((get_SV32_PTE_bits n_pte :: 32 Word.word)) Supervisor False False False) \ ((\ (w__2 :: bool MemoryOpResult) . (case w__2 of MemValue (_) => return () | MemException (e) => internal_error (''invalid physical address in TLB'') ) \ return (TR_Address ((or_vec(TLB_Entry_pAddr ent) ((EXTZ (( 34 :: int)::ii) ((and_vec vAddr(TLB_Entry_vAddrMask ent) :: 32 Word.word)) :: 34 Word.word)) :: 34 Word.word), ext_ptw))))))) ) )))))) | None => (walk32 vAddr ac priv mxr do_sum ptb level False ext_ptw :: ( (( 34 Word.word), SV32_PTE)PTW_Result) M) \ ((\ (w__6 :: (( 34 Word.word), SV32_PTE) PTW_Result) . (case w__6 of PTW_Failure ((f, ext_ptw)) => return (TR_Failure (f, ext_ptw)) | PTW_Success ((pAddr, pte, pteAddr, level, global1, ext_ptw)) => (case ((update_PTE_Bits ((Mk_PTE_Bits ((get_SV32_PTE_BITS pte :: 8 Word.word)))) ac ((zeros_implicit (( 10 :: int)::ii) :: 10 Word.word)) :: ((PTE_Bits * 10 Word.word))option)) of None => add_to_TLB32 asid vAddr pAddr pte pteAddr level global1 \ return (TR_Address (pAddr, ext_ptw)) | Some ((pbits, ext)) => if ((\ ((plat_enable_dirty_update () )))) then return (TR_Failure (PTW_PTE_Update () , ext_ptw)) else (let (w_pte :: SV32_PTE) = (update_SV32_PTE_BITS pte ((get_PTE_Bits_bits pbits :: 8 Word.word))) in mem_write_value_priv ((to_phys_addr pteAddr :: 32 Word.word)) (( 4 :: int)::ii) ((get_SV32_PTE_bits w_pte :: 32 Word.word)) Supervisor False False False \ ((\ (w__7 :: bool MemoryOpResult) . (case w__7 of MemValue (_) => add_to_TLB32 asid vAddr pAddr w_pte pteAddr level global1 \ return (TR_Address (pAddr, ext_ptw)) | MemException (e) => return (TR_Failure (PTW_Access () , ext_ptw)) )))) ) ))) ))))\ for asid :: "(9)Word.word " and ptb :: "(34)Word.word " and vAddr :: "(32)Word.word " and ac :: "(unit)AccessType " and priv :: " Privilege " and mxr :: " bool " and do_sum :: " bool " and level :: " int " and ext_ptw :: " unit " \ \\val init_vmem_sv32 : unit -> M unit\\ definition init_vmem_sv32 :: \ unit \((register_value),(unit),(exception))monad \ where \ init_vmem_sv32 _ = ( write_reg tlb32_ref None )\ \ \\val legalize_satp : Architecture -> mword ty32 -> mword ty32 -> mword ty32\\ definition legalize_satp :: \ Architecture \(32)Word.word \(32)Word.word \(32)Word.word \ where \ legalize_satp (a :: Architecture) (o1 :: xlenbits) (v :: xlenbits) = ( (legalize_satp32 a o1 v :: 32 Word.word))\ for a :: " Architecture " and o1 :: "(32)Word.word " and v :: "(32)Word.word " \ \\val translationMode : Privilege -> M SATPMode\\ definition translationMode :: \ Privilege \((register_value),(SATPMode),(exception))monad \ where \ translationMode priv = ( if (((priv = Machine))) then return Sbare else read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . (let arch = (architecture ((get_mstatus_SXL w__0 :: 2 Word.word))) in (case arch of Some (RV32) => (read_reg satp_ref :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let s = (Mk_Satp32 ((subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))) in return (if (((((get_Satp32_Mode s :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then Sbare else Sv32)))) | _ => internal_error (''unsupported address translation arch'') )))))\ for priv :: " Privilege " \ \\val translateAddr_priv : mword ty32 -> AccessType unit -> Privilege -> M (TR_Result (mword ty32) ExceptionType)\\ definition translateAddr_priv :: \(32)Word.word \(unit)AccessType \ Privilege \((register_value),((((32)Word.word),(ExceptionType))TR_Result),(exception))monad \ where \ translateAddr_priv vAddr ac effPriv = ( read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . (let (mxr :: bool) = (((get_Mstatus_MXR w__0 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)) in read_reg mstatus_ref \ ((\ (w__1 :: Mstatus) . (let (do_sum :: bool) = (((get_Mstatus_SUM w__1 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)) in translationMode effPriv \ ((\ (mode :: SATPMode) . (read_reg satp_ref :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . (let asid = ((curAsid32 w__2 :: 9 Word.word)) in (read_reg satp_ref :: ( 32 Word.word) M) \ ((\ (w__3 :: 32 Word.word) . (let ptb = ((curPTB32 w__3 :: 34 Word.word)) in (let (ext_ptw :: ext_ptw) = init_ext_ptw in (case mode of Sbare => return (TR_Address (vAddr, ext_ptw)) | Sv32 => (translate32 asid ptb vAddr ac effPriv mxr do_sum ((SV32_LEVELS - (( 1 :: int)::ii))) ext_ptw :: ( (( 34 Word.word), PTW_Error)TR_Result) M) \ ((\ (w__4 :: (( 34 Word.word), PTW_Error) TR_Result) . return ((case w__4 of TR_Address ((pa, ext_ptw)) => TR_Address ((to_phys_addr pa :: 32 Word.word), ext_ptw) | TR_Failure ((f, ext_ptw)) => TR_Failure (translationException ac f, ext_ptw) )))) | _ => (internal_error (''unsupported address translation scheme'') :: ( (( 32 Word.word), ExceptionType)TR_Result) M) )))))))))))))))))\ for vAddr :: "(32)Word.word " and ac :: "(unit)AccessType " and effPriv :: " Privilege " \ \\val translateAddr : mword ty32 -> AccessType unit -> M (TR_Result (mword ty32) ExceptionType)\\ definition translateAddr :: \(32)Word.word \(unit)AccessType \((register_value),((((32)Word.word),(ExceptionType))TR_Result),(exception))monad \ where \ translateAddr vAddr ac = ( read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . read_reg cur_privilege_ref \ ((\ (w__1 :: Privilege) . effectivePrivilege ac w__0 w__1 \ ((\ (w__2 :: Privilege) . (translateAddr_priv vAddr ac w__2 :: ( (( 32 Word.word), ExceptionType)TR_Result) M))))))))\ for vAddr :: "(32)Word.word " and ac :: "(unit)AccessType " \ \\val flush_TLB : maybe (mword ty32) -> maybe (mword ty32) -> M unit\\ definition flush_TLB :: \((32)Word.word)option \((32)Word.word)option \((register_value),(unit),(exception))monad \ where \ flush_TLB asid_xlen addr_xlen = ( (let (asid :: asid32 option) = ((case asid_xlen of None => None | Some (a) => Some ((subrange_vec_dec a (( 8 :: int)::ii) (( 0 :: int)::ii) :: 9 Word.word)) )) in flush_TLB32 asid addr_xlen))\ for asid_xlen :: "((32)Word.word)option " and addr_xlen :: "((32)Word.word)option " \ \\val init_vmem : unit -> M unit\\ definition init_vmem :: \ unit \((register_value),(unit),(exception))monad \ where \ init_vmem _ = ( init_vmem_sv32 () )\ \ \\val execute : ast -> M Retired\\ \ \\val assembly_forwards : ast -> M string\\ \ \\val assembly_backwards : string -> M ast\\ \ \\val assembly_forwards_matches : ast -> bool\\ \ \\val assembly_backwards_matches : string -> bool\\ \ \\val assembly_matches_prefix : string -> maybe ((ast * ii))\\ \ \\val encdec_forwards : ast -> M (mword ty32)\\ \ \\val encdec_backwards : mword ty32 -> M ast\\ \ \\val encdec_forwards_matches : ast -> M bool\\ \ \\val encdec_backwards_matches : mword ty32 -> M bool\\ \ \\val encdec_compressed_forwards : ast -> M (mword ty16)\\ \ \\val encdec_compressed_backwards : mword ty16 -> M ast\\ \ \\val encdec_compressed_forwards_matches : ast -> M bool\\ \ \\val encdec_compressed_backwards_matches : mword ty16 -> M bool\\ \ \\val encdec_uop_forwards : uop -> mword ty7\\ \ \\val encdec_uop_backwards : mword ty7 -> M uop\\ \ \\val encdec_uop_forwards_matches : uop -> bool\\ \ \\val encdec_uop_backwards_matches : mword ty7 -> bool\\ fun encdec_uop_forwards :: \ uop \(7)Word.word \ where \ encdec_uop_forwards RISCV_LUI = ( ( 0b0110111 :: 7 Word.word))\ |\ encdec_uop_forwards RISCV_AUIPC = ( ( 0b0010111 :: 7 Word.word))\ definition encdec_uop_backwards :: \(7)Word.word \((register_value),(uop),(exception))monad \ where \ encdec_uop_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b0110111 :: 7 Word.word)))) then return RISCV_LUI else if (((b__0 = ( 0b0010111 :: 7 Word.word)))) then return RISCV_AUIPC else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(7)Word.word " fun encdec_uop_forwards_matches :: \ uop \ bool \ where \ encdec_uop_forwards_matches RISCV_LUI = ( True )\ |\ encdec_uop_forwards_matches RISCV_AUIPC = ( True )\ definition encdec_uop_backwards_matches :: \(7)Word.word \ bool \ where \ encdec_uop_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b0110111 :: 7 Word.word)))) then True else if (((b__0 = ( 0b0010111 :: 7 Word.word)))) then True else False))\ for arg1 :: "(7)Word.word " \ \\val utype_mnemonic_forwards : uop -> string\\ \ \\val utype_mnemonic_backwards : string -> M uop\\ \ \\val utype_mnemonic_forwards_matches : uop -> bool\\ \ \\val utype_mnemonic_backwards_matches : string -> bool\\ \ \\val utype_mnemonic_matches_prefix : string -> maybe ((uop * ii))\\ fun utype_mnemonic_forwards :: \ uop \ string \ where \ utype_mnemonic_forwards RISCV_LUI = ( (''lui''))\ |\ utype_mnemonic_forwards RISCV_AUIPC = ( (''auipc''))\ definition utype_mnemonic_backwards :: \ string \((register_value),(uop),(exception))monad \ where \ utype_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''lui'')))) then return RISCV_LUI else if (((p00 = (''auipc'')))) then return RISCV_AUIPC else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun utype_mnemonic_forwards_matches :: \ uop \ bool \ where \ utype_mnemonic_forwards_matches RISCV_LUI = ( True )\ |\ utype_mnemonic_forwards_matches RISCV_AUIPC = ( True )\ definition utype_mnemonic_backwards_matches :: \ string \ bool \ where \ utype_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''lui'')))) then True else if (((p00 = (''auipc'')))) then True else False))\ for arg1 :: " string " \ \\val _s761_ : string -> maybe string\\ definition s761 :: \ string \(string)option \ where \ s761 s7620 = ( (let s7630 = s7620 in if ((string_startswith s7630 (''auipc''))) then (case ((string_drop s7630 ((string_length (''auipc''))))) of s1 => Some s1 ) else None))\ for s7620 :: " string " \ \\val _s757_ : string -> maybe string\\ definition s757 :: \ string \(string)option \ where \ s757 s7580 = ( (let s7590 = s7580 in if ((string_startswith s7590 (''lui''))) then (case ((string_drop s7590 ((string_length (''lui''))))) of s1 => Some s1 ) else None))\ for s7580 :: " string " definition utype_mnemonic_matches_prefix :: \ string \(uop*int)option \ where \ utype_mnemonic_matches_prefix arg1 = ( (let s7600 = arg1 in if ((case ((s757 s7600)) of Some (s1) => True | _ => False )) then (case s757 s7600 of (Some (s1)) => Some (RISCV_LUI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s761 s7600)) of Some (s1) => True | _ => False )) then (case s761 s7600 of (Some (s1)) => Some (RISCV_AUIPC, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val encdec_bop_forwards : bop -> mword ty3\\ \ \\val encdec_bop_backwards : mword ty3 -> M bop\\ \ \\val encdec_bop_forwards_matches : bop -> bool\\ \ \\val encdec_bop_backwards_matches : mword ty3 -> bool\\ fun encdec_bop_forwards :: \ bop \(3)Word.word \ where \ encdec_bop_forwards RISCV_BEQ = ( ( 0b000 :: 3 Word.word))\ |\ encdec_bop_forwards RISCV_BNE = ( ( 0b001 :: 3 Word.word))\ |\ encdec_bop_forwards RISCV_BLT = ( ( 0b100 :: 3 Word.word))\ |\ encdec_bop_forwards RISCV_BGE = ( ( 0b101 :: 3 Word.word))\ |\ encdec_bop_forwards RISCV_BLTU = ( ( 0b110 :: 3 Word.word))\ |\ encdec_bop_forwards RISCV_BGEU = ( ( 0b111 :: 3 Word.word))\ definition encdec_bop_backwards :: \(3)Word.word \((register_value),(bop),(exception))monad \ where \ encdec_bop_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then return RISCV_BEQ else if (((b__0 = ( 0b001 :: 3 Word.word)))) then return RISCV_BNE else if (((b__0 = ( 0b100 :: 3 Word.word)))) then return RISCV_BLT else if (((b__0 = ( 0b101 :: 3 Word.word)))) then return RISCV_BGE else if (((b__0 = ( 0b110 :: 3 Word.word)))) then return RISCV_BLTU else if (((b__0 = ( 0b111 :: 3 Word.word)))) then return RISCV_BGEU else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(3)Word.word " fun encdec_bop_forwards_matches :: \ bop \ bool \ where \ encdec_bop_forwards_matches RISCV_BEQ = ( True )\ |\ encdec_bop_forwards_matches RISCV_BNE = ( True )\ |\ encdec_bop_forwards_matches RISCV_BLT = ( True )\ |\ encdec_bop_forwards_matches RISCV_BGE = ( True )\ |\ encdec_bop_forwards_matches RISCV_BLTU = ( True )\ |\ encdec_bop_forwards_matches RISCV_BGEU = ( True )\ definition encdec_bop_backwards_matches :: \(3)Word.word \ bool \ where \ encdec_bop_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then True else if (((b__0 = ( 0b001 :: 3 Word.word)))) then True else if (((b__0 = ( 0b100 :: 3 Word.word)))) then True else if (((b__0 = ( 0b101 :: 3 Word.word)))) then True else if (((b__0 = ( 0b110 :: 3 Word.word)))) then True else if (((b__0 = ( 0b111 :: 3 Word.word)))) then True else False))\ for arg1 :: "(3)Word.word " \ \\val btype_mnemonic_forwards : bop -> string\\ \ \\val btype_mnemonic_backwards : string -> M bop\\ \ \\val btype_mnemonic_forwards_matches : bop -> bool\\ \ \\val btype_mnemonic_backwards_matches : string -> bool\\ \ \\val btype_mnemonic_matches_prefix : string -> maybe ((bop * ii))\\ fun btype_mnemonic_forwards :: \ bop \ string \ where \ btype_mnemonic_forwards RISCV_BEQ = ( (''beq''))\ |\ btype_mnemonic_forwards RISCV_BNE = ( (''bne''))\ |\ btype_mnemonic_forwards RISCV_BLT = ( (''blt''))\ |\ btype_mnemonic_forwards RISCV_BGE = ( (''bge''))\ |\ btype_mnemonic_forwards RISCV_BLTU = ( (''bltu''))\ |\ btype_mnemonic_forwards RISCV_BGEU = ( (''bgeu''))\ definition btype_mnemonic_backwards :: \ string \((register_value),(bop),(exception))monad \ where \ btype_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''beq'')))) then return RISCV_BEQ else if (((p00 = (''bne'')))) then return RISCV_BNE else if (((p00 = (''blt'')))) then return RISCV_BLT else if (((p00 = (''bge'')))) then return RISCV_BGE else if (((p00 = (''bltu'')))) then return RISCV_BLTU else if (((p00 = (''bgeu'')))) then return RISCV_BGEU else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun btype_mnemonic_forwards_matches :: \ bop \ bool \ where \ btype_mnemonic_forwards_matches RISCV_BEQ = ( True )\ |\ btype_mnemonic_forwards_matches RISCV_BNE = ( True )\ |\ btype_mnemonic_forwards_matches RISCV_BLT = ( True )\ |\ btype_mnemonic_forwards_matches RISCV_BGE = ( True )\ |\ btype_mnemonic_forwards_matches RISCV_BLTU = ( True )\ |\ btype_mnemonic_forwards_matches RISCV_BGEU = ( True )\ definition btype_mnemonic_backwards_matches :: \ string \ bool \ where \ btype_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''beq'')))) then True else if (((p00 = (''bne'')))) then True else if (((p00 = (''blt'')))) then True else if (((p00 = (''bge'')))) then True else if (((p00 = (''bltu'')))) then True else if (((p00 = (''bgeu'')))) then True else False))\ for arg1 :: " string " \ \\val _s785_ : string -> maybe string\\ definition s785 :: \ string \(string)option \ where \ s785 s7860 = ( (let s7870 = s7860 in if ((string_startswith s7870 (''bgeu''))) then (case ((string_drop s7870 ((string_length (''bgeu''))))) of s1 => Some s1 ) else None))\ for s7860 :: " string " \ \\val _s781_ : string -> maybe string\\ definition s781 :: \ string \(string)option \ where \ s781 s7820 = ( (let s7830 = s7820 in if ((string_startswith s7830 (''bltu''))) then (case ((string_drop s7830 ((string_length (''bltu''))))) of s1 => Some s1 ) else None))\ for s7820 :: " string " \ \\val _s777_ : string -> maybe string\\ definition s777 :: \ string \(string)option \ where \ s777 s7780 = ( (let s7790 = s7780 in if ((string_startswith s7790 (''bge''))) then (case ((string_drop s7790 ((string_length (''bge''))))) of s1 => Some s1 ) else None))\ for s7780 :: " string " \ \\val _s773_ : string -> maybe string\\ definition s773 :: \ string \(string)option \ where \ s773 s7740 = ( (let s7750 = s7740 in if ((string_startswith s7750 (''blt''))) then (case ((string_drop s7750 ((string_length (''blt''))))) of s1 => Some s1 ) else None))\ for s7740 :: " string " \ \\val _s769_ : string -> maybe string\\ definition s769 :: \ string \(string)option \ where \ s769 s7700 = ( (let s7710 = s7700 in if ((string_startswith s7710 (''bne''))) then (case ((string_drop s7710 ((string_length (''bne''))))) of s1 => Some s1 ) else None))\ for s7700 :: " string " \ \\val _s765_ : string -> maybe string\\ definition s765 :: \ string \(string)option \ where \ s765 s7660 = ( (let s7670 = s7660 in if ((string_startswith s7670 (''beq''))) then (case ((string_drop s7670 ((string_length (''beq''))))) of s1 => Some s1 ) else None))\ for s7660 :: " string " definition btype_mnemonic_matches_prefix :: \ string \(bop*int)option \ where \ btype_mnemonic_matches_prefix arg1 = ( (let s7680 = arg1 in if ((case ((s765 s7680)) of Some (s1) => True | _ => False )) then (case s765 s7680 of (Some (s1)) => Some (RISCV_BEQ, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s769 s7680)) of Some (s1) => True | _ => False )) then (case s769 s7680 of (Some (s1)) => Some (RISCV_BNE, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s773 s7680)) of Some (s1) => True | _ => False )) then (case s773 s7680 of (Some (s1)) => Some (RISCV_BLT, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s777 s7680)) of Some (s1) => True | _ => False )) then (case s777 s7680 of (Some (s1)) => Some (RISCV_BGE, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s781 s7680)) of Some (s1) => True | _ => False )) then (case s781 s7680 of (Some (s1)) => Some (RISCV_BLTU, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s785 s7680)) of Some (s1) => True | _ => False )) then (case s785 s7680 of (Some (s1)) => Some (RISCV_BGEU, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val encdec_iop_forwards : iop -> mword ty3\\ \ \\val encdec_iop_backwards : mword ty3 -> M iop\\ \ \\val encdec_iop_forwards_matches : iop -> bool\\ \ \\val encdec_iop_backwards_matches : mword ty3 -> bool\\ fun encdec_iop_forwards :: \ iop \(3)Word.word \ where \ encdec_iop_forwards RISCV_ADDI = ( ( 0b000 :: 3 Word.word))\ |\ encdec_iop_forwards RISCV_SLTI = ( ( 0b010 :: 3 Word.word))\ |\ encdec_iop_forwards RISCV_SLTIU = ( ( 0b011 :: 3 Word.word))\ |\ encdec_iop_forwards RISCV_ANDI = ( ( 0b111 :: 3 Word.word))\ |\ encdec_iop_forwards RISCV_ORI = ( ( 0b110 :: 3 Word.word))\ |\ encdec_iop_forwards RISCV_XORI = ( ( 0b100 :: 3 Word.word))\ definition encdec_iop_backwards :: \(3)Word.word \((register_value),(iop),(exception))monad \ where \ encdec_iop_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then return RISCV_ADDI else if (((b__0 = ( 0b010 :: 3 Word.word)))) then return RISCV_SLTI else if (((b__0 = ( 0b011 :: 3 Word.word)))) then return RISCV_SLTIU else if (((b__0 = ( 0b111 :: 3 Word.word)))) then return RISCV_ANDI else if (((b__0 = ( 0b110 :: 3 Word.word)))) then return RISCV_ORI else if (((b__0 = ( 0b100 :: 3 Word.word)))) then return RISCV_XORI else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(3)Word.word " fun encdec_iop_forwards_matches :: \ iop \ bool \ where \ encdec_iop_forwards_matches RISCV_ADDI = ( True )\ |\ encdec_iop_forwards_matches RISCV_SLTI = ( True )\ |\ encdec_iop_forwards_matches RISCV_SLTIU = ( True )\ |\ encdec_iop_forwards_matches RISCV_ANDI = ( True )\ |\ encdec_iop_forwards_matches RISCV_ORI = ( True )\ |\ encdec_iop_forwards_matches RISCV_XORI = ( True )\ definition encdec_iop_backwards_matches :: \(3)Word.word \ bool \ where \ encdec_iop_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then True else if (((b__0 = ( 0b010 :: 3 Word.word)))) then True else if (((b__0 = ( 0b011 :: 3 Word.word)))) then True else if (((b__0 = ( 0b111 :: 3 Word.word)))) then True else if (((b__0 = ( 0b110 :: 3 Word.word)))) then True else if (((b__0 = ( 0b100 :: 3 Word.word)))) then True else False))\ for arg1 :: "(3)Word.word " \ \\val itype_mnemonic_forwards : iop -> string\\ \ \\val itype_mnemonic_backwards : string -> M iop\\ \ \\val itype_mnemonic_forwards_matches : iop -> bool\\ \ \\val itype_mnemonic_backwards_matches : string -> bool\\ \ \\val itype_mnemonic_matches_prefix : string -> maybe ((iop * ii))\\ fun itype_mnemonic_forwards :: \ iop \ string \ where \ itype_mnemonic_forwards RISCV_ADDI = ( (''addi''))\ |\ itype_mnemonic_forwards RISCV_SLTI = ( (''slti''))\ |\ itype_mnemonic_forwards RISCV_SLTIU = ( (''sltiu''))\ |\ itype_mnemonic_forwards RISCV_XORI = ( (''xori''))\ |\ itype_mnemonic_forwards RISCV_ORI = ( (''ori''))\ |\ itype_mnemonic_forwards RISCV_ANDI = ( (''andi''))\ definition itype_mnemonic_backwards :: \ string \((register_value),(iop),(exception))monad \ where \ itype_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''addi'')))) then return RISCV_ADDI else if (((p00 = (''slti'')))) then return RISCV_SLTI else if (((p00 = (''sltiu'')))) then return RISCV_SLTIU else if (((p00 = (''xori'')))) then return RISCV_XORI else if (((p00 = (''ori'')))) then return RISCV_ORI else if (((p00 = (''andi'')))) then return RISCV_ANDI else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun itype_mnemonic_forwards_matches :: \ iop \ bool \ where \ itype_mnemonic_forwards_matches RISCV_ADDI = ( True )\ |\ itype_mnemonic_forwards_matches RISCV_SLTI = ( True )\ |\ itype_mnemonic_forwards_matches RISCV_SLTIU = ( True )\ |\ itype_mnemonic_forwards_matches RISCV_XORI = ( True )\ |\ itype_mnemonic_forwards_matches RISCV_ORI = ( True )\ |\ itype_mnemonic_forwards_matches RISCV_ANDI = ( True )\ definition itype_mnemonic_backwards_matches :: \ string \ bool \ where \ itype_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''addi'')))) then True else if (((p00 = (''slti'')))) then True else if (((p00 = (''sltiu'')))) then True else if (((p00 = (''xori'')))) then True else if (((p00 = (''ori'')))) then True else if (((p00 = (''andi'')))) then True else False))\ for arg1 :: " string " \ \\val _s809_ : string -> maybe string\\ definition s809 :: \ string \(string)option \ where \ s809 s8100 = ( (let s8110 = s8100 in if ((string_startswith s8110 (''andi''))) then (case ((string_drop s8110 ((string_length (''andi''))))) of s1 => Some s1 ) else None))\ for s8100 :: " string " \ \\val _s805_ : string -> maybe string\\ definition s805 :: \ string \(string)option \ where \ s805 s8060 = ( (let s8070 = s8060 in if ((string_startswith s8070 (''ori''))) then (case ((string_drop s8070 ((string_length (''ori''))))) of s1 => Some s1 ) else None))\ for s8060 :: " string " \ \\val _s801_ : string -> maybe string\\ definition s801 :: \ string \(string)option \ where \ s801 s8020 = ( (let s8030 = s8020 in if ((string_startswith s8030 (''xori''))) then (case ((string_drop s8030 ((string_length (''xori''))))) of s1 => Some s1 ) else None))\ for s8020 :: " string " \ \\val _s797_ : string -> maybe string\\ definition s797 :: \ string \(string)option \ where \ s797 s7980 = ( (let s7990 = s7980 in if ((string_startswith s7990 (''sltiu''))) then (case ((string_drop s7990 ((string_length (''sltiu''))))) of s1 => Some s1 ) else None))\ for s7980 :: " string " \ \\val _s793_ : string -> maybe string\\ definition s793 :: \ string \(string)option \ where \ s793 s7940 = ( (let s7950 = s7940 in if ((string_startswith s7950 (''slti''))) then (case ((string_drop s7950 ((string_length (''slti''))))) of s1 => Some s1 ) else None))\ for s7940 :: " string " \ \\val _s789_ : string -> maybe string\\ definition s789 :: \ string \(string)option \ where \ s789 s7900 = ( (let s7910 = s7900 in if ((string_startswith s7910 (''addi''))) then (case ((string_drop s7910 ((string_length (''addi''))))) of s1 => Some s1 ) else None))\ for s7900 :: " string " definition itype_mnemonic_matches_prefix :: \ string \(iop*int)option \ where \ itype_mnemonic_matches_prefix arg1 = ( (let s7920 = arg1 in if ((case ((s789 s7920)) of Some (s1) => True | _ => False )) then (case s789 s7920 of (Some (s1)) => Some (RISCV_ADDI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s793 s7920)) of Some (s1) => True | _ => False )) then (case s793 s7920 of (Some (s1)) => Some (RISCV_SLTI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s797 s7920)) of Some (s1) => True | _ => False )) then (case s797 s7920 of (Some (s1)) => Some (RISCV_SLTIU, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s801 s7920)) of Some (s1) => True | _ => False )) then (case s801 s7920 of (Some (s1)) => Some (RISCV_XORI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s805 s7920)) of Some (s1) => True | _ => False )) then (case s805 s7920 of (Some (s1)) => Some (RISCV_ORI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s809 s7920)) of Some (s1) => True | _ => False )) then (case s809 s7920 of (Some (s1)) => Some (RISCV_ANDI, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val encdec_sop_forwards : sop -> mword ty3\\ \ \\val encdec_sop_backwards : mword ty3 -> M sop\\ \ \\val encdec_sop_forwards_matches : sop -> bool\\ \ \\val encdec_sop_backwards_matches : mword ty3 -> bool\\ fun encdec_sop_forwards :: \ sop \(3)Word.word \ where \ encdec_sop_forwards RISCV_SLLI = ( ( 0b001 :: 3 Word.word))\ |\ encdec_sop_forwards RISCV_SRLI = ( ( 0b101 :: 3 Word.word))\ |\ encdec_sop_forwards RISCV_SRAI = ( ( 0b101 :: 3 Word.word))\ definition encdec_sop_backwards :: \(3)Word.word \((register_value),(sop),(exception))monad \ where \ encdec_sop_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b001 :: 3 Word.word)))) then return RISCV_SLLI else if (((b__0 = ( 0b101 :: 3 Word.word)))) then return RISCV_SRLI else if (((b__0 = ( 0b101 :: 3 Word.word)))) then return RISCV_SRAI else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(3)Word.word " fun encdec_sop_forwards_matches :: \ sop \ bool \ where \ encdec_sop_forwards_matches RISCV_SLLI = ( True )\ |\ encdec_sop_forwards_matches RISCV_SRLI = ( True )\ |\ encdec_sop_forwards_matches RISCV_SRAI = ( True )\ definition encdec_sop_backwards_matches :: \(3)Word.word \ bool \ where \ encdec_sop_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b001 :: 3 Word.word)))) then True else if (((b__0 = ( 0b101 :: 3 Word.word)))) then True else if (((b__0 = ( 0b101 :: 3 Word.word)))) then True else False))\ for arg1 :: "(3)Word.word " \ \\val shiftiop_mnemonic_forwards : sop -> string\\ \ \\val shiftiop_mnemonic_backwards : string -> M sop\\ \ \\val shiftiop_mnemonic_forwards_matches : sop -> bool\\ \ \\val shiftiop_mnemonic_backwards_matches : string -> bool\\ \ \\val shiftiop_mnemonic_matches_prefix : string -> maybe ((sop * ii))\\ fun shiftiop_mnemonic_forwards :: \ sop \ string \ where \ shiftiop_mnemonic_forwards RISCV_SLLI = ( (''slli''))\ |\ shiftiop_mnemonic_forwards RISCV_SRLI = ( (''srli''))\ |\ shiftiop_mnemonic_forwards RISCV_SRAI = ( (''srai''))\ definition shiftiop_mnemonic_backwards :: \ string \((register_value),(sop),(exception))monad \ where \ shiftiop_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''slli'')))) then return RISCV_SLLI else if (((p00 = (''srli'')))) then return RISCV_SRLI else if (((p00 = (''srai'')))) then return RISCV_SRAI else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun shiftiop_mnemonic_forwards_matches :: \ sop \ bool \ where \ shiftiop_mnemonic_forwards_matches RISCV_SLLI = ( True )\ |\ shiftiop_mnemonic_forwards_matches RISCV_SRLI = ( True )\ |\ shiftiop_mnemonic_forwards_matches RISCV_SRAI = ( True )\ definition shiftiop_mnemonic_backwards_matches :: \ string \ bool \ where \ shiftiop_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''slli'')))) then True else if (((p00 = (''srli'')))) then True else if (((p00 = (''srai'')))) then True else False))\ for arg1 :: " string " \ \\val _s821_ : string -> maybe string\\ definition s821 :: \ string \(string)option \ where \ s821 s8220 = ( (let s8230 = s8220 in if ((string_startswith s8230 (''srai''))) then (case ((string_drop s8230 ((string_length (''srai''))))) of s1 => Some s1 ) else None))\ for s8220 :: " string " \ \\val _s817_ : string -> maybe string\\ definition s817 :: \ string \(string)option \ where \ s817 s8180 = ( (let s8190 = s8180 in if ((string_startswith s8190 (''srli''))) then (case ((string_drop s8190 ((string_length (''srli''))))) of s1 => Some s1 ) else None))\ for s8180 :: " string " \ \\val _s813_ : string -> maybe string\\ definition s813 :: \ string \(string)option \ where \ s813 s8140 = ( (let s8150 = s8140 in if ((string_startswith s8150 (''slli''))) then (case ((string_drop s8150 ((string_length (''slli''))))) of s1 => Some s1 ) else None))\ for s8140 :: " string " definition shiftiop_mnemonic_matches_prefix :: \ string \(sop*int)option \ where \ shiftiop_mnemonic_matches_prefix arg1 = ( (let s8160 = arg1 in if ((case ((s813 s8160)) of Some (s1) => True | _ => False )) then (case s813 s8160 of (Some (s1)) => Some (RISCV_SLLI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s817 s8160)) of Some (s1) => True | _ => False )) then (case s817 s8160 of (Some (s1)) => Some (RISCV_SRLI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s821 s8160)) of Some (s1) => True | _ => False )) then (case s821 s8160 of (Some (s1)) => Some (RISCV_SRAI, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val rtype_mnemonic_forwards : rop -> string\\ \ \\val rtype_mnemonic_backwards : string -> M rop\\ \ \\val rtype_mnemonic_forwards_matches : rop -> bool\\ \ \\val rtype_mnemonic_backwards_matches : string -> bool\\ \ \\val rtype_mnemonic_matches_prefix : string -> maybe ((rop * ii))\\ fun rtype_mnemonic_forwards :: \ rop \ string \ where \ rtype_mnemonic_forwards RISCV_ADD = ( (''add''))\ |\ rtype_mnemonic_forwards RISCV_SLT = ( (''slt''))\ |\ rtype_mnemonic_forwards RISCV_SLTU = ( (''sltu''))\ |\ rtype_mnemonic_forwards RISCV_AND = ( (''and''))\ |\ rtype_mnemonic_forwards RISCV_OR = ( (''or''))\ |\ rtype_mnemonic_forwards RISCV_XOR = ( (''xor''))\ |\ rtype_mnemonic_forwards RISCV_SLL = ( (''sll''))\ |\ rtype_mnemonic_forwards RISCV_SRL = ( (''srl''))\ |\ rtype_mnemonic_forwards RISCV_SUB = ( (''sub''))\ |\ rtype_mnemonic_forwards RISCV_SRA = ( (''sra''))\ definition rtype_mnemonic_backwards :: \ string \((register_value),(rop),(exception))monad \ where \ rtype_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''add'')))) then return RISCV_ADD else if (((p00 = (''slt'')))) then return RISCV_SLT else if (((p00 = (''sltu'')))) then return RISCV_SLTU else if (((p00 = (''and'')))) then return RISCV_AND else if (((p00 = (''or'')))) then return RISCV_OR else if (((p00 = (''xor'')))) then return RISCV_XOR else if (((p00 = (''sll'')))) then return RISCV_SLL else if (((p00 = (''srl'')))) then return RISCV_SRL else if (((p00 = (''sub'')))) then return RISCV_SUB else if (((p00 = (''sra'')))) then return RISCV_SRA else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun rtype_mnemonic_forwards_matches :: \ rop \ bool \ where \ rtype_mnemonic_forwards_matches RISCV_ADD = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_SLT = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_SLTU = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_AND = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_OR = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_XOR = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_SLL = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_SRL = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_SUB = ( True )\ |\ rtype_mnemonic_forwards_matches RISCV_SRA = ( True )\ definition rtype_mnemonic_backwards_matches :: \ string \ bool \ where \ rtype_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''add'')))) then True else if (((p00 = (''slt'')))) then True else if (((p00 = (''sltu'')))) then True else if (((p00 = (''and'')))) then True else if (((p00 = (''or'')))) then True else if (((p00 = (''xor'')))) then True else if (((p00 = (''sll'')))) then True else if (((p00 = (''srl'')))) then True else if (((p00 = (''sub'')))) then True else if (((p00 = (''sra'')))) then True else False))\ for arg1 :: " string " \ \\val _s861_ : string -> maybe string\\ definition s861 :: \ string \(string)option \ where \ s861 s8620 = ( (let s8630 = s8620 in if ((string_startswith s8630 (''sra''))) then (case ((string_drop s8630 ((string_length (''sra''))))) of s1 => Some s1 ) else None))\ for s8620 :: " string " \ \\val _s857_ : string -> maybe string\\ definition s857 :: \ string \(string)option \ where \ s857 s8580 = ( (let s8590 = s8580 in if ((string_startswith s8590 (''sub''))) then (case ((string_drop s8590 ((string_length (''sub''))))) of s1 => Some s1 ) else None))\ for s8580 :: " string " \ \\val _s853_ : string -> maybe string\\ definition s853 :: \ string \(string)option \ where \ s853 s8540 = ( (let s8550 = s8540 in if ((string_startswith s8550 (''srl''))) then (case ((string_drop s8550 ((string_length (''srl''))))) of s1 => Some s1 ) else None))\ for s8540 :: " string " \ \\val _s849_ : string -> maybe string\\ definition s849 :: \ string \(string)option \ where \ s849 s8500 = ( (let s8510 = s8500 in if ((string_startswith s8510 (''sll''))) then (case ((string_drop s8510 ((string_length (''sll''))))) of s1 => Some s1 ) else None))\ for s8500 :: " string " \ \\val _s845_ : string -> maybe string\\ definition s845 :: \ string \(string)option \ where \ s845 s8460 = ( (let s8470 = s8460 in if ((string_startswith s8470 (''xor''))) then (case ((string_drop s8470 ((string_length (''xor''))))) of s1 => Some s1 ) else None))\ for s8460 :: " string " \ \\val _s841_ : string -> maybe string\\ definition s841 :: \ string \(string)option \ where \ s841 s8420 = ( (let s8430 = s8420 in if ((string_startswith s8430 (''or''))) then (case ((string_drop s8430 ((string_length (''or''))))) of s1 => Some s1 ) else None))\ for s8420 :: " string " \ \\val _s837_ : string -> maybe string\\ definition s837 :: \ string \(string)option \ where \ s837 s8380 = ( (let s8390 = s8380 in if ((string_startswith s8390 (''and''))) then (case ((string_drop s8390 ((string_length (''and''))))) of s1 => Some s1 ) else None))\ for s8380 :: " string " \ \\val _s833_ : string -> maybe string\\ definition s833 :: \ string \(string)option \ where \ s833 s8340 = ( (let s8350 = s8340 in if ((string_startswith s8350 (''sltu''))) then (case ((string_drop s8350 ((string_length (''sltu''))))) of s1 => Some s1 ) else None))\ for s8340 :: " string " \ \\val _s829_ : string -> maybe string\\ definition s829 :: \ string \(string)option \ where \ s829 s8300 = ( (let s8310 = s8300 in if ((string_startswith s8310 (''slt''))) then (case ((string_drop s8310 ((string_length (''slt''))))) of s1 => Some s1 ) else None))\ for s8300 :: " string " \ \\val _s825_ : string -> maybe string\\ definition s825 :: \ string \(string)option \ where \ s825 s8260 = ( (let s8270 = s8260 in if ((string_startswith s8270 (''add''))) then (case ((string_drop s8270 ((string_length (''add''))))) of s1 => Some s1 ) else None))\ for s8260 :: " string " definition rtype_mnemonic_matches_prefix :: \ string \(rop*int)option \ where \ rtype_mnemonic_matches_prefix arg1 = ( (let s8280 = arg1 in if ((case ((s825 s8280)) of Some (s1) => True | _ => False )) then (case s825 s8280 of (Some (s1)) => Some (RISCV_ADD, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s829 s8280)) of Some (s1) => True | _ => False )) then (case s829 s8280 of (Some (s1)) => Some (RISCV_SLT, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s833 s8280)) of Some (s1) => True | _ => False )) then (case s833 s8280 of (Some (s1)) => Some (RISCV_SLTU, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s837 s8280)) of Some (s1) => True | _ => False )) then (case s837 s8280 of (Some (s1)) => Some (RISCV_AND, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s841 s8280)) of Some (s1) => True | _ => False )) then (case s841 s8280 of (Some (s1)) => Some (RISCV_OR, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s845 s8280)) of Some (s1) => True | _ => False )) then (case s845 s8280 of (Some (s1)) => Some (RISCV_XOR, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s849 s8280)) of Some (s1) => True | _ => False )) then (case s849 s8280 of (Some (s1)) => Some (RISCV_SLL, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s853 s8280)) of Some (s1) => True | _ => False )) then (case s853 s8280 of (Some (s1)) => Some (RISCV_SRL, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s857 s8280)) of Some (s1) => True | _ => False )) then (case s857 s8280 of (Some (s1)) => Some (RISCV_SUB, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s861 s8280)) of Some (s1) => True | _ => False )) then (case s861 s8280 of (Some (s1)) => Some (RISCV_SRA, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val extend_value : forall 'int8_times_n. Size 'int8_times_n => bool -> MemoryOpResult (mword 'int8_times_n) -> MemoryOpResult (mword ty32)\\ fun extend_value :: \ bool \(('int8_times_n::len)Word.word)MemoryOpResult \((32)Word.word)MemoryOpResult \ where \ extend_value is_unsigned (MemValue (v)) = ( MemValue (if is_unsigned then (EXTZ (( 32 :: int)::ii) v :: 32 Word.word) else (EXTS (( 32 :: int)::ii) v :: 32 Word.word)))\ for is_unsigned :: " bool " and v :: "('int8_times_n::len)Word.word " |\ extend_value is_unsigned (MemException (e)) = ( MemException e )\ for is_unsigned :: " bool " and e :: " ExceptionType " \ \\val process_load : forall 'int8_times_n. Size 'int8_times_n => mword ty5 -> mword ty32 -> MemoryOpResult (mword 'int8_times_n) -> bool -> M Retired\\ definition process_load :: \(5)Word.word \(32)Word.word \(('int8_times_n::len)Word.word)MemoryOpResult \ bool \((register_value),(Retired),(exception))monad \ where \ process_load rd vaddr value1 is_unsigned = ( (case ((extend_value is_unsigned value1 :: ( 32 Word.word) MemoryOpResult)) of MemValue (result) => wX_bits rd result \ return RETIRE_SUCCESS | MemException (e) => handle_mem_exception vaddr e \ return RETIRE_FAIL ))\ for rd :: "(5)Word.word " and vaddr :: "(32)Word.word " and value1 :: "(('int8_times_n::len)Word.word)MemoryOpResult " and is_unsigned :: " bool " \ \\val check_misaligned : mword ty32 -> word_width -> bool\\ definition check_misaligned :: \(32)Word.word \ word_width \ bool \ where \ check_misaligned (vaddr :: xlenbits) (width :: word_width) = ( if ((plat_enable_misaligned_access () )) then False else (case width of BYTE => False | HALF => (((access_vec_dec vaddr (( 0 :: int)::ii))) = B1) | WORD => ((((((access_vec_dec vaddr (( 0 :: int)::ii))) = B1))) \ (((((access_vec_dec vaddr (( 1 :: int)::ii))) = B1)))) | DOUBLE => ((((((access_vec_dec vaddr (( 0 :: int)::ii))) = B1))) \ ((((((((access_vec_dec vaddr (( 1 :: int)::ii))) = B1))) \ (((((access_vec_dec vaddr (( 2 :: int)::ii))) = B1))))))) ))\ for vaddr :: "(32)Word.word " and width :: " word_width " \ \\val maybe_aq_forwards : bool -> string\\ \ \\val maybe_aq_backwards : string -> M bool\\ \ \\val maybe_aq_forwards_matches : bool -> bool\\ \ \\val maybe_aq_backwards_matches : string -> bool\\ \ \\val maybe_aq_matches_prefix : string -> maybe ((bool * ii))\\ fun maybe_aq_forwards :: \ bool \ string \ where \ maybe_aq_forwards True = ( (''.aq''))\ |\ maybe_aq_forwards False = ( (''''))\ definition maybe_aq_backwards :: \ string \((register_value),(bool),(exception))monad \ where \ maybe_aq_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''.aq'')))) then return True else if (((p00 = ('''')))) then return False else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun maybe_aq_forwards_matches :: \ bool \ bool \ where \ maybe_aq_forwards_matches True = ( True )\ |\ maybe_aq_forwards_matches False = ( True )\ definition maybe_aq_backwards_matches :: \ string \ bool \ where \ maybe_aq_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''.aq'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s869_ : string -> maybe string\\ definition s869 :: \ string \(string)option \ where \ s869 s8700 = ( (let s8710 = s8700 in if ((string_startswith s8710 (''''))) then (case ((string_drop s8710 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s8700 :: " string " \ \\val _s865_ : string -> maybe string\\ definition s865 :: \ string \(string)option \ where \ s865 s8660 = ( (let s8670 = s8660 in if ((string_startswith s8670 (''.aq''))) then (case ((string_drop s8670 ((string_length (''.aq''))))) of s1 => Some s1 ) else None))\ for s8660 :: " string " definition maybe_aq_matches_prefix :: \ string \(bool*int)option \ where \ maybe_aq_matches_prefix arg1 = ( (let s8680 = arg1 in if ((case ((s865 s8680)) of Some (s1) => True | _ => False )) then (case s865 s8680 of (Some (s1)) => Some (True, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s869 s8680)) of Some (s1) => True | _ => False )) then (case s869 s8680 of (Some (s1)) => Some (False, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val maybe_rl_forwards : bool -> string\\ \ \\val maybe_rl_backwards : string -> M bool\\ \ \\val maybe_rl_forwards_matches : bool -> bool\\ \ \\val maybe_rl_backwards_matches : string -> bool\\ \ \\val maybe_rl_matches_prefix : string -> maybe ((bool * ii))\\ fun maybe_rl_forwards :: \ bool \ string \ where \ maybe_rl_forwards True = ( (''.rl''))\ |\ maybe_rl_forwards False = ( (''''))\ definition maybe_rl_backwards :: \ string \((register_value),(bool),(exception))monad \ where \ maybe_rl_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''.rl'')))) then return True else if (((p00 = ('''')))) then return False else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun maybe_rl_forwards_matches :: \ bool \ bool \ where \ maybe_rl_forwards_matches True = ( True )\ |\ maybe_rl_forwards_matches False = ( True )\ definition maybe_rl_backwards_matches :: \ string \ bool \ where \ maybe_rl_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''.rl'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s877_ : string -> maybe string\\ definition s877 :: \ string \(string)option \ where \ s877 s8780 = ( (let s8790 = s8780 in if ((string_startswith s8790 (''''))) then (case ((string_drop s8790 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s8780 :: " string " \ \\val _s873_ : string -> maybe string\\ definition s873 :: \ string \(string)option \ where \ s873 s8740 = ( (let s8750 = s8740 in if ((string_startswith s8750 (''.rl''))) then (case ((string_drop s8750 ((string_length (''.rl''))))) of s1 => Some s1 ) else None))\ for s8740 :: " string " definition maybe_rl_matches_prefix :: \ string \(bool*int)option \ where \ maybe_rl_matches_prefix arg1 = ( (let s8760 = arg1 in if ((case ((s873 s8760)) of Some (s1) => True | _ => False )) then (case s873 s8760 of (Some (s1)) => Some (True, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s877 s8760)) of Some (s1) => True | _ => False )) then (case s877 s8760 of (Some (s1)) => Some (False, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val maybe_u_forwards : bool -> string\\ \ \\val maybe_u_backwards : string -> M bool\\ \ \\val maybe_u_forwards_matches : bool -> bool\\ \ \\val maybe_u_backwards_matches : string -> bool\\ \ \\val maybe_u_matches_prefix : string -> maybe ((bool * ii))\\ fun maybe_u_forwards :: \ bool \ string \ where \ maybe_u_forwards True = ( (''u''))\ |\ maybe_u_forwards False = ( (''''))\ definition maybe_u_backwards :: \ string \((register_value),(bool),(exception))monad \ where \ maybe_u_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''u'')))) then return True else if (((p00 = ('''')))) then return False else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun maybe_u_forwards_matches :: \ bool \ bool \ where \ maybe_u_forwards_matches True = ( True )\ |\ maybe_u_forwards_matches False = ( True )\ definition maybe_u_backwards_matches :: \ string \ bool \ where \ maybe_u_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''u'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s885_ : string -> maybe string\\ definition s885 :: \ string \(string)option \ where \ s885 s8860 = ( (let s8870 = s8860 in if ((string_startswith s8870 (''''))) then (case ((string_drop s8870 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s8860 :: " string " \ \\val _s881_ : string -> maybe string\\ definition s881 :: \ string \(string)option \ where \ s881 s8820 = ( (let s8830 = s8820 in if ((string_startswith s8830 (''u''))) then (case ((string_drop s8830 ((string_length (''u''))))) of s1 => Some s1 ) else None))\ for s8820 :: " string " definition maybe_u_matches_prefix :: \ string \(bool*int)option \ where \ maybe_u_matches_prefix arg1 = ( (let s8840 = arg1 in if ((case ((s881 s8840)) of Some (s1) => True | _ => False )) then (case s881 s8840 of (Some (s1)) => Some (True, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s885 s8840)) of Some (s1) => True | _ => False )) then (case s885 s8840 of (Some (s1)) => Some (False, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val shiftw_mnemonic_forwards : sop -> string\\ \ \\val shiftw_mnemonic_backwards : string -> M sop\\ \ \\val shiftw_mnemonic_forwards_matches : sop -> bool\\ \ \\val shiftw_mnemonic_backwards_matches : string -> bool\\ \ \\val shiftw_mnemonic_matches_prefix : string -> maybe ((sop * ii))\\ fun shiftw_mnemonic_forwards :: \ sop \ string \ where \ shiftw_mnemonic_forwards RISCV_SLLI = ( (''slli''))\ |\ shiftw_mnemonic_forwards RISCV_SRLI = ( (''srli''))\ |\ shiftw_mnemonic_forwards RISCV_SRAI = ( (''srai''))\ definition shiftw_mnemonic_backwards :: \ string \((register_value),(sop),(exception))monad \ where \ shiftw_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''slli'')))) then return RISCV_SLLI else if (((p00 = (''srli'')))) then return RISCV_SRLI else if (((p00 = (''srai'')))) then return RISCV_SRAI else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun shiftw_mnemonic_forwards_matches :: \ sop \ bool \ where \ shiftw_mnemonic_forwards_matches RISCV_SLLI = ( True )\ |\ shiftw_mnemonic_forwards_matches RISCV_SRLI = ( True )\ |\ shiftw_mnemonic_forwards_matches RISCV_SRAI = ( True )\ definition shiftw_mnemonic_backwards_matches :: \ string \ bool \ where \ shiftw_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''slli'')))) then True else if (((p00 = (''srli'')))) then True else if (((p00 = (''srai'')))) then True else False))\ for arg1 :: " string " \ \\val _s897_ : string -> maybe string\\ definition s897 :: \ string \(string)option \ where \ s897 s8980 = ( (let s8990 = s8980 in if ((string_startswith s8990 (''srai''))) then (case ((string_drop s8990 ((string_length (''srai''))))) of s1 => Some s1 ) else None))\ for s8980 :: " string " \ \\val _s893_ : string -> maybe string\\ definition s893 :: \ string \(string)option \ where \ s893 s8940 = ( (let s8950 = s8940 in if ((string_startswith s8950 (''srli''))) then (case ((string_drop s8950 ((string_length (''srli''))))) of s1 => Some s1 ) else None))\ for s8940 :: " string " \ \\val _s889_ : string -> maybe string\\ definition s889 :: \ string \(string)option \ where \ s889 s8900 = ( (let s8910 = s8900 in if ((string_startswith s8910 (''slli''))) then (case ((string_drop s8910 ((string_length (''slli''))))) of s1 => Some s1 ) else None))\ for s8900 :: " string " definition shiftw_mnemonic_matches_prefix :: \ string \(sop*int)option \ where \ shiftw_mnemonic_matches_prefix arg1 = ( (let s8920 = arg1 in if ((case ((s889 s8920)) of Some (s1) => True | _ => False )) then (case s889 s8920 of (Some (s1)) => Some (RISCV_SLLI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s893 s8920)) of Some (s1) => True | _ => False )) then (case s893 s8920 of (Some (s1)) => Some (RISCV_SRLI, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s897 s8920)) of Some (s1) => True | _ => False )) then (case s897 s8920 of (Some (s1)) => Some (RISCV_SRAI, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val rtypew_mnemonic_forwards : ropw -> string\\ \ \\val rtypew_mnemonic_backwards : string -> M ropw\\ \ \\val rtypew_mnemonic_forwards_matches : ropw -> bool\\ \ \\val rtypew_mnemonic_backwards_matches : string -> bool\\ \ \\val rtypew_mnemonic_matches_prefix : string -> maybe ((ropw * ii))\\ fun rtypew_mnemonic_forwards :: \ ropw \ string \ where \ rtypew_mnemonic_forwards RISCV_ADDW = ( (''addw''))\ |\ rtypew_mnemonic_forwards RISCV_SUBW = ( (''subw''))\ |\ rtypew_mnemonic_forwards RISCV_SLLW = ( (''sllw''))\ |\ rtypew_mnemonic_forwards RISCV_SRLW = ( (''srlw''))\ |\ rtypew_mnemonic_forwards RISCV_SRAW = ( (''sraw''))\ definition rtypew_mnemonic_backwards :: \ string \((register_value),(ropw),(exception))monad \ where \ rtypew_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''addw'')))) then return RISCV_ADDW else if (((p00 = (''subw'')))) then return RISCV_SUBW else if (((p00 = (''sllw'')))) then return RISCV_SLLW else if (((p00 = (''srlw'')))) then return RISCV_SRLW else if (((p00 = (''sraw'')))) then return RISCV_SRAW else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun rtypew_mnemonic_forwards_matches :: \ ropw \ bool \ where \ rtypew_mnemonic_forwards_matches RISCV_ADDW = ( True )\ |\ rtypew_mnemonic_forwards_matches RISCV_SUBW = ( True )\ |\ rtypew_mnemonic_forwards_matches RISCV_SLLW = ( True )\ |\ rtypew_mnemonic_forwards_matches RISCV_SRLW = ( True )\ |\ rtypew_mnemonic_forwards_matches RISCV_SRAW = ( True )\ definition rtypew_mnemonic_backwards_matches :: \ string \ bool \ where \ rtypew_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''addw'')))) then True else if (((p00 = (''subw'')))) then True else if (((p00 = (''sllw'')))) then True else if (((p00 = (''srlw'')))) then True else if (((p00 = (''sraw'')))) then True else False))\ for arg1 :: " string " \ \\val _s917_ : string -> maybe string\\ definition s917 :: \ string \(string)option \ where \ s917 s9180 = ( (let s9190 = s9180 in if ((string_startswith s9190 (''sraw''))) then (case ((string_drop s9190 ((string_length (''sraw''))))) of s1 => Some s1 ) else None))\ for s9180 :: " string " \ \\val _s913_ : string -> maybe string\\ definition s913 :: \ string \(string)option \ where \ s913 s9140 = ( (let s9150 = s9140 in if ((string_startswith s9150 (''srlw''))) then (case ((string_drop s9150 ((string_length (''srlw''))))) of s1 => Some s1 ) else None))\ for s9140 :: " string " \ \\val _s909_ : string -> maybe string\\ definition s909 :: \ string \(string)option \ where \ s909 s9100 = ( (let s9110 = s9100 in if ((string_startswith s9110 (''sllw''))) then (case ((string_drop s9110 ((string_length (''sllw''))))) of s1 => Some s1 ) else None))\ for s9100 :: " string " \ \\val _s905_ : string -> maybe string\\ definition s905 :: \ string \(string)option \ where \ s905 s9060 = ( (let s9070 = s9060 in if ((string_startswith s9070 (''subw''))) then (case ((string_drop s9070 ((string_length (''subw''))))) of s1 => Some s1 ) else None))\ for s9060 :: " string " \ \\val _s901_ : string -> maybe string\\ definition s901 :: \ string \(string)option \ where \ s901 s9020 = ( (let s9030 = s9020 in if ((string_startswith s9030 (''addw''))) then (case ((string_drop s9030 ((string_length (''addw''))))) of s1 => Some s1 ) else None))\ for s9020 :: " string " definition rtypew_mnemonic_matches_prefix :: \ string \(ropw*int)option \ where \ rtypew_mnemonic_matches_prefix arg1 = ( (let s9040 = arg1 in if ((case ((s901 s9040)) of Some (s1) => True | _ => False )) then (case s901 s9040 of (Some (s1)) => Some (RISCV_ADDW, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s905 s9040)) of Some (s1) => True | _ => False )) then (case s905 s9040 of (Some (s1)) => Some (RISCV_SUBW, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s909 s9040)) of Some (s1) => True | _ => False )) then (case s909 s9040 of (Some (s1)) => Some (RISCV_SLLW, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s913 s9040)) of Some (s1) => True | _ => False )) then (case s913 s9040 of (Some (s1)) => Some (RISCV_SRLW, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s917 s9040)) of Some (s1) => True | _ => False )) then (case s917 s9040 of (Some (s1)) => Some (RISCV_SRAW, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val shiftiwop_mnemonic_forwards : sopw -> string\\ \ \\val shiftiwop_mnemonic_backwards : string -> M sopw\\ \ \\val shiftiwop_mnemonic_forwards_matches : sopw -> bool\\ \ \\val shiftiwop_mnemonic_backwards_matches : string -> bool\\ \ \\val shiftiwop_mnemonic_matches_prefix : string -> maybe ((sopw * ii))\\ fun shiftiwop_mnemonic_forwards :: \ sopw \ string \ where \ shiftiwop_mnemonic_forwards RISCV_SLLIW = ( (''slliw''))\ |\ shiftiwop_mnemonic_forwards RISCV_SRLIW = ( (''srliw''))\ |\ shiftiwop_mnemonic_forwards RISCV_SRAIW = ( (''sraiw''))\ definition shiftiwop_mnemonic_backwards :: \ string \((register_value),(sopw),(exception))monad \ where \ shiftiwop_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''slliw'')))) then return RISCV_SLLIW else if (((p00 = (''srliw'')))) then return RISCV_SRLIW else if (((p00 = (''sraiw'')))) then return RISCV_SRAIW else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun shiftiwop_mnemonic_forwards_matches :: \ sopw \ bool \ where \ shiftiwop_mnemonic_forwards_matches RISCV_SLLIW = ( True )\ |\ shiftiwop_mnemonic_forwards_matches RISCV_SRLIW = ( True )\ |\ shiftiwop_mnemonic_forwards_matches RISCV_SRAIW = ( True )\ definition shiftiwop_mnemonic_backwards_matches :: \ string \ bool \ where \ shiftiwop_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''slliw'')))) then True else if (((p00 = (''srliw'')))) then True else if (((p00 = (''sraiw'')))) then True else False))\ for arg1 :: " string " \ \\val _s929_ : string -> maybe string\\ definition s929 :: \ string \(string)option \ where \ s929 s9300 = ( (let s9310 = s9300 in if ((string_startswith s9310 (''sraiw''))) then (case ((string_drop s9310 ((string_length (''sraiw''))))) of s1 => Some s1 ) else None))\ for s9300 :: " string " \ \\val _s925_ : string -> maybe string\\ definition s925 :: \ string \(string)option \ where \ s925 s9260 = ( (let s9270 = s9260 in if ((string_startswith s9270 (''srliw''))) then (case ((string_drop s9270 ((string_length (''srliw''))))) of s1 => Some s1 ) else None))\ for s9260 :: " string " \ \\val _s921_ : string -> maybe string\\ definition s921 :: \ string \(string)option \ where \ s921 s9220 = ( (let s9230 = s9220 in if ((string_startswith s9230 (''slliw''))) then (case ((string_drop s9230 ((string_length (''slliw''))))) of s1 => Some s1 ) else None))\ for s9220 :: " string " definition shiftiwop_mnemonic_matches_prefix :: \ string \(sopw*int)option \ where \ shiftiwop_mnemonic_matches_prefix arg1 = ( (let s9240 = arg1 in if ((case ((s921 s9240)) of Some (s1) => True | _ => False )) then (case s921 s9240 of (Some (s1)) => Some (RISCV_SLLIW, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s925 s9240)) of Some (s1) => True | _ => False )) then (case s925 s9240 of (Some (s1)) => Some (RISCV_SRLIW, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s929 s9240)) of Some (s1) => True | _ => False )) then (case s929 s9240 of (Some (s1)) => Some (RISCV_SRAIW, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val bit_maybe_r_forwards : mword ty1 -> M string\\ \ \\val bit_maybe_r_backwards : string -> M (mword ty1)\\ \ \\val bit_maybe_r_forwards_matches : mword ty1 -> bool\\ \ \\val bit_maybe_r_backwards_matches : string -> bool\\ \ \\val bit_maybe_r_matches_prefix : string -> maybe ((mword ty1 * ii))\\ definition bit_maybe_r_forwards :: \(1)Word.word \((register_value),(string),(exception))monad \ where \ bit_maybe_r_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then return (''r'') else if (((b__0 = ( 0b0 :: 1 Word.word)))) then return ('''') else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(1)Word.word " definition bit_maybe_r_backwards :: \ string \((register_value),((1)Word.word),(exception))monad \ where \ bit_maybe_r_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''r'')))) then return ( 0b1 :: 1 Word.word) else if (((p00 = ('''')))) then return ( 0b0 :: 1 Word.word) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition bit_maybe_r_forwards_matches :: \(1)Word.word \ bool \ where \ bit_maybe_r_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then True else if (((b__0 = ( 0b0 :: 1 Word.word)))) then True else False))\ for arg1 :: "(1)Word.word " definition bit_maybe_r_backwards_matches :: \ string \ bool \ where \ bit_maybe_r_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''r'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s937_ : string -> maybe string\\ definition s937 :: \ string \(string)option \ where \ s937 s9380 = ( (let s9390 = s9380 in if ((string_startswith s9390 (''''))) then (case ((string_drop s9390 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s9380 :: " string " \ \\val _s933_ : string -> maybe string\\ definition s933 :: \ string \(string)option \ where \ s933 s9340 = ( (let s9350 = s9340 in if ((string_startswith s9350 (''r''))) then (case ((string_drop s9350 ((string_length (''r''))))) of s1 => Some s1 ) else None))\ for s9340 :: " string " definition bit_maybe_r_matches_prefix :: \ string \((1)Word.word*int)option \ where \ bit_maybe_r_matches_prefix arg1 = ( (let s9360 = arg1 in if ((case ((s933 s9360)) of Some (s1) => True | _ => False )) then (case s933 s9360 of (Some (s1)) => Some (( 0b1 :: 1 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s937 s9360)) of Some (s1) => True | _ => False )) then (case s937 s9360 of (Some (s1)) => Some (( 0b0 :: 1 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val bit_maybe_w_forwards : mword ty1 -> M string\\ \ \\val bit_maybe_w_backwards : string -> M (mword ty1)\\ \ \\val bit_maybe_w_forwards_matches : mword ty1 -> bool\\ \ \\val bit_maybe_w_backwards_matches : string -> bool\\ \ \\val bit_maybe_w_matches_prefix : string -> maybe ((mword ty1 * ii))\\ definition bit_maybe_w_forwards :: \(1)Word.word \((register_value),(string),(exception))monad \ where \ bit_maybe_w_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then return (''w'') else if (((b__0 = ( 0b0 :: 1 Word.word)))) then return ('''') else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(1)Word.word " definition bit_maybe_w_backwards :: \ string \((register_value),((1)Word.word),(exception))monad \ where \ bit_maybe_w_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''w'')))) then return ( 0b1 :: 1 Word.word) else if (((p00 = ('''')))) then return ( 0b0 :: 1 Word.word) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition bit_maybe_w_forwards_matches :: \(1)Word.word \ bool \ where \ bit_maybe_w_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then True else if (((b__0 = ( 0b0 :: 1 Word.word)))) then True else False))\ for arg1 :: "(1)Word.word " definition bit_maybe_w_backwards_matches :: \ string \ bool \ where \ bit_maybe_w_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''w'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s945_ : string -> maybe string\\ definition s945 :: \ string \(string)option \ where \ s945 s9460 = ( (let s9470 = s9460 in if ((string_startswith s9470 (''''))) then (case ((string_drop s9470 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s9460 :: " string " \ \\val _s941_ : string -> maybe string\\ definition s941 :: \ string \(string)option \ where \ s941 s9420 = ( (let s9430 = s9420 in if ((string_startswith s9430 (''w''))) then (case ((string_drop s9430 ((string_length (''w''))))) of s1 => Some s1 ) else None))\ for s9420 :: " string " definition bit_maybe_w_matches_prefix :: \ string \((1)Word.word*int)option \ where \ bit_maybe_w_matches_prefix arg1 = ( (let s9440 = arg1 in if ((case ((s941 s9440)) of Some (s1) => True | _ => False )) then (case s941 s9440 of (Some (s1)) => Some (( 0b1 :: 1 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s945 s9440)) of Some (s1) => True | _ => False )) then (case s945 s9440 of (Some (s1)) => Some (( 0b0 :: 1 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val bit_maybe_i_forwards : mword ty1 -> M string\\ \ \\val bit_maybe_i_backwards : string -> M (mword ty1)\\ \ \\val bit_maybe_i_forwards_matches : mword ty1 -> bool\\ \ \\val bit_maybe_i_backwards_matches : string -> bool\\ \ \\val bit_maybe_i_matches_prefix : string -> maybe ((mword ty1 * ii))\\ definition bit_maybe_i_forwards :: \(1)Word.word \((register_value),(string),(exception))monad \ where \ bit_maybe_i_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then return (''i'') else if (((b__0 = ( 0b0 :: 1 Word.word)))) then return ('''') else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(1)Word.word " definition bit_maybe_i_backwards :: \ string \((register_value),((1)Word.word),(exception))monad \ where \ bit_maybe_i_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''i'')))) then return ( 0b1 :: 1 Word.word) else if (((p00 = ('''')))) then return ( 0b0 :: 1 Word.word) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition bit_maybe_i_forwards_matches :: \(1)Word.word \ bool \ where \ bit_maybe_i_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then True else if (((b__0 = ( 0b0 :: 1 Word.word)))) then True else False))\ for arg1 :: "(1)Word.word " definition bit_maybe_i_backwards_matches :: \ string \ bool \ where \ bit_maybe_i_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''i'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s953_ : string -> maybe string\\ definition s953 :: \ string \(string)option \ where \ s953 s9540 = ( (let s9550 = s9540 in if ((string_startswith s9550 (''''))) then (case ((string_drop s9550 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s9540 :: " string " \ \\val _s949_ : string -> maybe string\\ definition s949 :: \ string \(string)option \ where \ s949 s9500 = ( (let s9510 = s9500 in if ((string_startswith s9510 (''i''))) then (case ((string_drop s9510 ((string_length (''i''))))) of s1 => Some s1 ) else None))\ for s9500 :: " string " definition bit_maybe_i_matches_prefix :: \ string \((1)Word.word*int)option \ where \ bit_maybe_i_matches_prefix arg1 = ( (let s9520 = arg1 in if ((case ((s949 s9520)) of Some (s1) => True | _ => False )) then (case s949 s9520 of (Some (s1)) => Some (( 0b1 :: 1 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s953 s9520)) of Some (s1) => True | _ => False )) then (case s953 s9520 of (Some (s1)) => Some (( 0b0 :: 1 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val bit_maybe_o_forwards : mword ty1 -> M string\\ \ \\val bit_maybe_o_backwards : string -> M (mword ty1)\\ \ \\val bit_maybe_o_forwards_matches : mword ty1 -> bool\\ \ \\val bit_maybe_o_backwards_matches : string -> bool\\ \ \\val bit_maybe_o_matches_prefix : string -> maybe ((mword ty1 * ii))\\ definition bit_maybe_o_forwards :: \(1)Word.word \((register_value),(string),(exception))monad \ where \ bit_maybe_o_forwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then return (''o'') else if (((b__0 = ( 0b0 :: 1 Word.word)))) then return ('''') else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(1)Word.word " definition bit_maybe_o_backwards :: \ string \((register_value),((1)Word.word),(exception))monad \ where \ bit_maybe_o_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''o'')))) then return ( 0b1 :: 1 Word.word) else if (((p00 = ('''')))) then return ( 0b0 :: 1 Word.word) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition bit_maybe_o_forwards_matches :: \(1)Word.word \ bool \ where \ bit_maybe_o_forwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b1 :: 1 Word.word)))) then True else if (((b__0 = ( 0b0 :: 1 Word.word)))) then True else False))\ for arg1 :: "(1)Word.word " definition bit_maybe_o_backwards_matches :: \ string \ bool \ where \ bit_maybe_o_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''o'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s961_ : string -> maybe string\\ definition s961 :: \ string \(string)option \ where \ s961 s9620 = ( (let s9630 = s9620 in if ((string_startswith s9630 (''''))) then (case ((string_drop s9630 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s9620 :: " string " \ \\val _s957_ : string -> maybe string\\ definition s957 :: \ string \(string)option \ where \ s957 s9580 = ( (let s9590 = s9580 in if ((string_startswith s9590 (''o''))) then (case ((string_drop s9590 ((string_length (''o''))))) of s1 => Some s1 ) else None))\ for s9580 :: " string " definition bit_maybe_o_matches_prefix :: \ string \((1)Word.word*int)option \ where \ bit_maybe_o_matches_prefix arg1 = ( (let s9600 = arg1 in if ((case ((s957 s9600)) of Some (s1) => True | _ => False )) then (case s957 s9600 of (Some (s1)) => Some (( 0b1 :: 1 Word.word), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s961 s9600)) of Some (s1) => True | _ => False )) then (case s961 s9600 of (Some (s1)) => Some (( 0b0 :: 1 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val fence_bits_forwards : mword ty4 -> M string\\ \ \\val fence_bits_backwards : string -> M (mword ty4)\\ \ \\val fence_bits_forwards_matches : mword ty4 -> bool\\ \ \\val fence_bits_backwards_matches : string -> bool\\ \ \\val fence_bits_matches_prefix : string -> maybe ((mword ty4 * ii))\\ definition fence_bits_forwards :: \(4)Word.word \((register_value),(string),(exception))monad \ where \ fence_bits_forwards v__0 = ( (let (i :: 1 bits) = ((subrange_vec_dec v__0 (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word)) in (let (w :: 1 bits) = ((subrange_vec_dec v__0 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in (let (r :: 1 bits) = ((subrange_vec_dec v__0 (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word)) in (let (o1 :: 1 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (i :: 1 bits) = ((subrange_vec_dec v__0 (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word)) in bit_maybe_i_forwards i \ ((\ (w__0 :: string) . bit_maybe_o_forwards o1 \ ((\ (w__1 :: string) . bit_maybe_r_forwards r \ ((\ (w__2 :: string) . bit_maybe_w_forwards w \ ((\ (w__3 :: string) . return ((string_append w__0 ((string_append w__1 ((string_append w__2 ((string_append w__3 ('''')))))))))))))))))))))))\ for v__0 :: "(4)Word.word " \ \\val _s965_ : string -> maybe ((mword ty1 * mword ty1 * mword ty1 * mword ty1))\\ definition s965 :: \ string \((1)Word.word*(1)Word.word*(1)Word.word*(1)Word.word)option \ where \ s965 s9670 = ( (case ((bit_maybe_i_matches_prefix s9670 :: (( 1 Word.word * ii))option)) of Some ((i, s9680)) => (case ((string_drop s9670 s9680)) of s9690 => (case ((bit_maybe_o_matches_prefix s9690 :: (( 1 Word.word * ii)) option)) of Some ((o1, s9700)) => (case ((string_drop s9690 s9700)) of s9710 => (case ((bit_maybe_r_matches_prefix s9710 :: (( 1 Word.word * ii)) option)) of Some ((r, s9720)) => (case ((string_drop s9710 s9720)) of s9730 => (case ((bit_maybe_w_matches_prefix s9730 :: (( 1 Word.word * ii)) option)) of Some ((w, s9740)) => (let p00 = (string_drop s9730 s9740) in if (((p00 = ('''')))) then Some (i, o1, r, w) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s9670 :: " string " definition fence_bits_backwards :: \ string \((register_value),((4)Word.word),(exception))monad \ where \ fence_bits_backwards arg1 = ( (let s9750 = arg1 in if ((case ((s965 s9750 :: (( 1 Word.word * 1 Word.word * 1 Word.word * 1 Word.word))option)) of Some ((i, o1, r, w)) => True | _ => False )) then (case (s965 s9750 :: (( 1 Word.word * 1 Word.word * 1 Word.word * 1 Word.word)) option) of (Some ((i, o1, r, w))) => return ((concat_vec i ((concat_vec o1 ((concat_vec r w :: 2 Word.word)) :: 3 Word.word)) :: 4 Word.word)) ) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition fence_bits_forwards_matches :: \(4)Word.word \ bool \ where \ fence_bits_forwards_matches v__1 = ( True )\ for v__1 :: "(4)Word.word " \ \\val _s976_ : string -> maybe ((mword ty1 * mword ty1 * mword ty1 * mword ty1))\\ definition s976 :: \ string \((1)Word.word*(1)Word.word*(1)Word.word*(1)Word.word)option \ where \ s976 s9780 = ( (case ((bit_maybe_i_matches_prefix s9780 :: (( 1 Word.word * ii))option)) of Some ((i, s9790)) => (case ((string_drop s9780 s9790)) of s9800 => (case ((bit_maybe_o_matches_prefix s9800 :: (( 1 Word.word * ii)) option)) of Some ((o1, s9810)) => (case ((string_drop s9800 s9810)) of s9820 => (case ((bit_maybe_r_matches_prefix s9820 :: (( 1 Word.word * ii)) option)) of Some ((r, s9830)) => (case ((string_drop s9820 s9830)) of s9840 => (case ((bit_maybe_w_matches_prefix s9840 :: (( 1 Word.word * ii)) option)) of Some ((w, s9850)) => (let p00 = (string_drop s9840 s9850) in if (((p00 = ('''')))) then Some (i, o1, r, w) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s9780 :: " string " definition fence_bits_backwards_matches :: \ string \ bool \ where \ fence_bits_backwards_matches arg1 = ( (let s9860 = arg1 in if ((case ((s976 s9860 :: (( 1 Word.word * 1 Word.word * 1 Word.word * 1 Word.word))option)) of Some ((i, o1, r, w)) => True | _ => False )) then (case (s976 s9860 :: (( 1 Word.word * 1 Word.word * 1 Word.word * 1 Word.word)) option) of (Some ((i, o1, r, w))) => True ) else False))\ for arg1 :: " string " \ \\val _s987_ : string -> maybe ((mword ty1 * mword ty1 * mword ty1 * mword ty1 * string))\\ definition s987 :: \ string \((1)Word.word*(1)Word.word*(1)Word.word*(1)Word.word*string)option \ where \ s987 s9890 = ( (case ((bit_maybe_i_matches_prefix s9890 :: (( 1 Word.word * ii))option)) of Some ((i, s9900)) => (case ((string_drop s9890 s9900)) of s9910 => (case ((bit_maybe_o_matches_prefix s9910 :: (( 1 Word.word * ii)) option)) of Some ((o1, s9920)) => (case ((string_drop s9910 s9920)) of s9930 => (case ((bit_maybe_r_matches_prefix s9930 :: (( 1 Word.word * ii)) option)) of Some ((r, s9940)) => (case ((string_drop s9930 s9940)) of s9950 => (case ((bit_maybe_w_matches_prefix s9950 :: (( 1 Word.word * ii)) option)) of Some ((w, s9960)) => (case ((string_drop s9950 s9960)) of s1 => Some (i, o1, r, w, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s9890 :: " string " definition fence_bits_matches_prefix :: \ string \((4)Word.word*int)option \ where \ fence_bits_matches_prefix arg1 = ( (let s9970 = arg1 in if ((case ((s987 s9970 :: (( 1 Word.word * 1 Word.word * 1 Word.word * 1 Word.word * string))option)) of Some ((i, o1, r, w, s1)) => True | _ => False )) then (case (s987 s9970 :: (( 1 Word.word * 1 Word.word * 1 Word.word * 1 Word.word * string)) option) of (Some ((i, o1, r, w, s1))) => Some ((concat_vec i ((concat_vec o1 ((concat_vec r w :: 2 Word.word)) :: 3 Word.word)) :: 4 Word.word), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val aqrl_str : bool -> bool -> string\\ fun aqrl_str :: \ bool \ bool \ string \ where \ aqrl_str (False :: bool) (False :: bool) = ( (''''))\ |\ aqrl_str (False :: bool) (True :: bool) = ( (''.rl''))\ |\ aqrl_str (True :: bool) (False :: bool) = ( (''.aq''))\ |\ aqrl_str (True :: bool) (True :: bool) = ( (''.aqrl''))\ \ \\val lrsc_width_str : word_width -> string\\ fun lrsc_width_str :: \ word_width \ string \ where \ lrsc_width_str BYTE = ( (''.b''))\ |\ lrsc_width_str HALF = ( (''.h''))\ |\ lrsc_width_str WORD = ( (''.w''))\ |\ lrsc_width_str DOUBLE = ( (''.d''))\ \ \\val amo_width_valid : word_width -> bool\\ fun amo_width_valid :: \ word_width \ bool \ where \ amo_width_valid WORD = ( True )\ |\ amo_width_valid DOUBLE = ( (( 32 :: int)::ii) \ (( 64 :: int)::ii))\ |\ amo_width_valid _ = ( False )\ \ \\val process_loadres : forall 'int8_times_n. Size 'int8_times_n => mword ty5 -> mword ty32 -> MemoryOpResult (mword 'int8_times_n) -> bool -> M Retired\\ definition process_loadres :: \(5)Word.word \(32)Word.word \(('int8_times_n::len)Word.word)MemoryOpResult \ bool \((register_value),(Retired),(exception))monad \ where \ process_loadres rd addr value1 is_unsigned = ( (case ((extend_value is_unsigned value1 :: ( 32 Word.word) MemoryOpResult)) of MemValue (result) => (let (_ :: unit) = (load_reservation addr) in wX_bits rd result \ return RETIRE_SUCCESS) | MemException (e) => handle_mem_exception addr e \ return RETIRE_FAIL ))\ for rd :: "(5)Word.word " and addr :: "(32)Word.word " and value1 :: "(('int8_times_n::len)Word.word)MemoryOpResult " and is_unsigned :: " bool " \ \\val encdec_amoop_forwards : amoop -> mword ty5\\ \ \\val encdec_amoop_backwards : mword ty5 -> M amoop\\ \ \\val encdec_amoop_forwards_matches : amoop -> bool\\ \ \\val encdec_amoop_backwards_matches : mword ty5 -> bool\\ fun encdec_amoop_forwards :: \ amoop \(5)Word.word \ where \ encdec_amoop_forwards AMOSWAP = ( ( 0b00001 :: 5 Word.word))\ |\ encdec_amoop_forwards AMOADD = ( ( 0b00000 :: 5 Word.word))\ |\ encdec_amoop_forwards AMOXOR = ( ( 0b00100 :: 5 Word.word))\ |\ encdec_amoop_forwards AMOAND = ( ( 0b01100 :: 5 Word.word))\ |\ encdec_amoop_forwards AMOOR = ( ( 0b01000 :: 5 Word.word))\ |\ encdec_amoop_forwards AMOMIN = ( ( 0b10000 :: 5 Word.word))\ |\ encdec_amoop_forwards AMOMAX = ( ( 0b10100 :: 5 Word.word))\ |\ encdec_amoop_forwards AMOMINU = ( ( 0b11000 :: 5 Word.word))\ |\ encdec_amoop_forwards AMOMAXU = ( ( 0b11100 :: 5 Word.word))\ definition encdec_amoop_backwards :: \(5)Word.word \((register_value),(amoop),(exception))monad \ where \ encdec_amoop_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00001 :: 5 Word.word)))) then return AMOSWAP else if (((b__0 = ( 0b00000 :: 5 Word.word)))) then return AMOADD else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then return AMOXOR else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then return AMOAND else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then return AMOOR else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then return AMOMIN else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then return AMOMAX else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then return AMOMINU else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then return AMOMAXU else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(5)Word.word " fun encdec_amoop_forwards_matches :: \ amoop \ bool \ where \ encdec_amoop_forwards_matches AMOSWAP = ( True )\ |\ encdec_amoop_forwards_matches AMOADD = ( True )\ |\ encdec_amoop_forwards_matches AMOXOR = ( True )\ |\ encdec_amoop_forwards_matches AMOAND = ( True )\ |\ encdec_amoop_forwards_matches AMOOR = ( True )\ |\ encdec_amoop_forwards_matches AMOMIN = ( True )\ |\ encdec_amoop_forwards_matches AMOMAX = ( True )\ |\ encdec_amoop_forwards_matches AMOMINU = ( True )\ |\ encdec_amoop_forwards_matches AMOMAXU = ( True )\ definition encdec_amoop_backwards_matches :: \(5)Word.word \ bool \ where \ encdec_amoop_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b00001 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b00100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b01000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b10100 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11000 :: 5 Word.word)))) then True else if (((b__0 = ( 0b11100 :: 5 Word.word)))) then True else False))\ for arg1 :: "(5)Word.word " \ \\val amo_mnemonic_forwards : amoop -> string\\ \ \\val amo_mnemonic_backwards : string -> M amoop\\ \ \\val amo_mnemonic_forwards_matches : amoop -> bool\\ \ \\val amo_mnemonic_backwards_matches : string -> bool\\ \ \\val amo_mnemonic_matches_prefix : string -> maybe ((amoop * ii))\\ fun amo_mnemonic_forwards :: \ amoop \ string \ where \ amo_mnemonic_forwards AMOSWAP = ( (''amoswap''))\ |\ amo_mnemonic_forwards AMOADD = ( (''amoadd''))\ |\ amo_mnemonic_forwards AMOXOR = ( (''amoxor''))\ |\ amo_mnemonic_forwards AMOAND = ( (''amoand''))\ |\ amo_mnemonic_forwards AMOOR = ( (''amoor''))\ |\ amo_mnemonic_forwards AMOMIN = ( (''amomin''))\ |\ amo_mnemonic_forwards AMOMAX = ( (''amomax''))\ |\ amo_mnemonic_forwards AMOMINU = ( (''amominu''))\ |\ amo_mnemonic_forwards AMOMAXU = ( (''amomaxu''))\ definition amo_mnemonic_backwards :: \ string \((register_value),(amoop),(exception))monad \ where \ amo_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''amoswap'')))) then return AMOSWAP else if (((p00 = (''amoadd'')))) then return AMOADD else if (((p00 = (''amoxor'')))) then return AMOXOR else if (((p00 = (''amoand'')))) then return AMOAND else if (((p00 = (''amoor'')))) then return AMOOR else if (((p00 = (''amomin'')))) then return AMOMIN else if (((p00 = (''amomax'')))) then return AMOMAX else if (((p00 = (''amominu'')))) then return AMOMINU else if (((p00 = (''amomaxu'')))) then return AMOMAXU else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun amo_mnemonic_forwards_matches :: \ amoop \ bool \ where \ amo_mnemonic_forwards_matches AMOSWAP = ( True )\ |\ amo_mnemonic_forwards_matches AMOADD = ( True )\ |\ amo_mnemonic_forwards_matches AMOXOR = ( True )\ |\ amo_mnemonic_forwards_matches AMOAND = ( True )\ |\ amo_mnemonic_forwards_matches AMOOR = ( True )\ |\ amo_mnemonic_forwards_matches AMOMIN = ( True )\ |\ amo_mnemonic_forwards_matches AMOMAX = ( True )\ |\ amo_mnemonic_forwards_matches AMOMINU = ( True )\ |\ amo_mnemonic_forwards_matches AMOMAXU = ( True )\ definition amo_mnemonic_backwards_matches :: \ string \ bool \ where \ amo_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''amoswap'')))) then True else if (((p00 = (''amoadd'')))) then True else if (((p00 = (''amoxor'')))) then True else if (((p00 = (''amoand'')))) then True else if (((p00 = (''amoor'')))) then True else if (((p00 = (''amomin'')))) then True else if (((p00 = (''amomax'')))) then True else if (((p00 = (''amominu'')))) then True else if (((p00 = (''amomaxu'')))) then True else False))\ for arg1 :: " string " \ \\val _s1030_ : string -> maybe string\\ definition s1030 :: \ string \(string)option \ where \ s1030 s10310 = ( (let s10320 = s10310 in if ((string_startswith s10320 (''amomaxu''))) then (case ((string_drop s10320 ((string_length (''amomaxu''))))) of s1 => Some s1 ) else None))\ for s10310 :: " string " \ \\val _s1026_ : string -> maybe string\\ definition s1026 :: \ string \(string)option \ where \ s1026 s10270 = ( (let s10280 = s10270 in if ((string_startswith s10280 (''amominu''))) then (case ((string_drop s10280 ((string_length (''amominu''))))) of s1 => Some s1 ) else None))\ for s10270 :: " string " \ \\val _s1022_ : string -> maybe string\\ definition s1022 :: \ string \(string)option \ where \ s1022 s10230 = ( (let s10240 = s10230 in if ((string_startswith s10240 (''amomax''))) then (case ((string_drop s10240 ((string_length (''amomax''))))) of s1 => Some s1 ) else None))\ for s10230 :: " string " \ \\val _s1018_ : string -> maybe string\\ definition s1018 :: \ string \(string)option \ where \ s1018 s10190 = ( (let s10200 = s10190 in if ((string_startswith s10200 (''amomin''))) then (case ((string_drop s10200 ((string_length (''amomin''))))) of s1 => Some s1 ) else None))\ for s10190 :: " string " \ \\val _s1014_ : string -> maybe string\\ definition s1014 :: \ string \(string)option \ where \ s1014 s10150 = ( (let s10160 = s10150 in if ((string_startswith s10160 (''amoor''))) then (case ((string_drop s10160 ((string_length (''amoor''))))) of s1 => Some s1 ) else None))\ for s10150 :: " string " \ \\val _s1010_ : string -> maybe string\\ definition s1010 :: \ string \(string)option \ where \ s1010 s10110 = ( (let s10120 = s10110 in if ((string_startswith s10120 (''amoand''))) then (case ((string_drop s10120 ((string_length (''amoand''))))) of s1 => Some s1 ) else None))\ for s10110 :: " string " \ \\val _s1006_ : string -> maybe string\\ definition s1006 :: \ string \(string)option \ where \ s1006 s10070 = ( (let s10080 = s10070 in if ((string_startswith s10080 (''amoxor''))) then (case ((string_drop s10080 ((string_length (''amoxor''))))) of s1 => Some s1 ) else None))\ for s10070 :: " string " \ \\val _s1002_ : string -> maybe string\\ definition s1002 :: \ string \(string)option \ where \ s1002 s10030 = ( (let s10040 = s10030 in if ((string_startswith s10040 (''amoadd''))) then (case ((string_drop s10040 ((string_length (''amoadd''))))) of s1 => Some s1 ) else None))\ for s10030 :: " string " \ \\val _s998_ : string -> maybe string\\ definition s998 :: \ string \(string)option \ where \ s998 s9990 = ( (let s10000 = s9990 in if ((string_startswith s10000 (''amoswap''))) then (case ((string_drop s10000 ((string_length (''amoswap''))))) of s1 => Some s1 ) else None))\ for s9990 :: " string " definition amo_mnemonic_matches_prefix :: \ string \(amoop*int)option \ where \ amo_mnemonic_matches_prefix arg1 = ( (let s10010 = arg1 in if ((case ((s998 s10010)) of Some (s1) => True | _ => False )) then (case s998 s10010 of (Some (s1)) => Some (AMOSWAP, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1002 s10010)) of Some (s1) => True | _ => False )) then (case s1002 s10010 of (Some (s1)) => Some (AMOADD, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1006 s10010)) of Some (s1) => True | _ => False )) then (case s1006 s10010 of (Some (s1)) => Some (AMOXOR, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1010 s10010)) of Some (s1) => True | _ => False )) then (case s1010 s10010 of (Some (s1)) => Some (AMOAND, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1014 s10010)) of Some (s1) => True | _ => False )) then (case s1014 s10010 of (Some (s1)) => Some (AMOOR, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1018 s10010)) of Some (s1) => True | _ => False )) then (case s1018 s10010 of (Some (s1)) => Some (AMOMIN, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1022 s10010)) of Some (s1) => True | _ => False )) then (case s1022 s10010 of (Some (s1)) => Some (AMOMAX, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1026 s10010)) of Some (s1) => True | _ => False )) then (case s1026 s10010 of (Some (s1)) => Some (AMOMINU, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1030 s10010)) of Some (s1) => True | _ => False )) then (case s1030 s10010 of (Some (s1)) => Some (AMOMAXU, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val encdec_mul_op_forwards : (bool * bool * bool) -> mword ty3\\ \ \\val encdec_mul_op_backwards : mword ty3 -> M (bool * bool * bool)\\ \ \\val encdec_mul_op_forwards_matches : (bool * bool * bool) -> bool\\ \ \\val encdec_mul_op_backwards_matches : mword ty3 -> bool\\ fun encdec_mul_op_forwards :: \ bool*bool*bool \(3)Word.word \ where \ encdec_mul_op_forwards (False, True, True) = ( ( 0b000 :: 3 Word.word))\ |\ encdec_mul_op_forwards (True, True, True) = ( ( 0b001 :: 3 Word.word))\ |\ encdec_mul_op_forwards (True, True, False) = ( ( 0b010 :: 3 Word.word))\ |\ encdec_mul_op_forwards (True, False, False) = ( ( 0b011 :: 3 Word.word))\ definition encdec_mul_op_backwards :: \(3)Word.word \((register_value),(bool*bool*bool),(exception))monad \ where \ encdec_mul_op_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then return (False, True, True) else if (((b__0 = ( 0b001 :: 3 Word.word)))) then return (True, True, True) else if (((b__0 = ( 0b010 :: 3 Word.word)))) then return (True, True, False) else if (((b__0 = ( 0b011 :: 3 Word.word)))) then return (True, False, False) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(3)Word.word " fun encdec_mul_op_forwards_matches :: \ bool*bool*bool \ bool \ where \ encdec_mul_op_forwards_matches (False, True, True) = ( True )\ |\ encdec_mul_op_forwards_matches (True, True, True) = ( True )\ |\ encdec_mul_op_forwards_matches (True, True, False) = ( True )\ |\ encdec_mul_op_forwards_matches (True, False, False) = ( True )\ |\ encdec_mul_op_forwards_matches _ = ( False )\ definition encdec_mul_op_backwards_matches :: \(3)Word.word \ bool \ where \ encdec_mul_op_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then True else if (((b__0 = ( 0b001 :: 3 Word.word)))) then True else if (((b__0 = ( 0b010 :: 3 Word.word)))) then True else if (((b__0 = ( 0b011 :: 3 Word.word)))) then True else False))\ for arg1 :: "(3)Word.word " \ \\val mul_mnemonic_forwards : (bool * bool * bool) -> string\\ \ \\val mul_mnemonic_backwards : string -> M (bool * bool * bool)\\ \ \\val mul_mnemonic_forwards_matches : (bool * bool * bool) -> bool\\ \ \\val mul_mnemonic_backwards_matches : string -> bool\\ \ \\val mul_mnemonic_matches_prefix : string -> maybe (((bool * bool * bool) * ii))\\ fun mul_mnemonic_forwards :: \ bool*bool*bool \ string \ where \ mul_mnemonic_forwards (False, True, True) = ( (''mul''))\ |\ mul_mnemonic_forwards (True, True, True) = ( (''mulh''))\ |\ mul_mnemonic_forwards (True, True, False) = ( (''mulhsu''))\ |\ mul_mnemonic_forwards (True, False, False) = ( (''mulhu''))\ definition mul_mnemonic_backwards :: \ string \((register_value),(bool*bool*bool),(exception))monad \ where \ mul_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''mul'')))) then return (False, True, True) else if (((p00 = (''mulh'')))) then return (True, True, True) else if (((p00 = (''mulhsu'')))) then return (True, True, False) else if (((p00 = (''mulhu'')))) then return (True, False, False) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun mul_mnemonic_forwards_matches :: \ bool*bool*bool \ bool \ where \ mul_mnemonic_forwards_matches (False, True, True) = ( True )\ |\ mul_mnemonic_forwards_matches (True, True, True) = ( True )\ |\ mul_mnemonic_forwards_matches (True, True, False) = ( True )\ |\ mul_mnemonic_forwards_matches (True, False, False) = ( True )\ |\ mul_mnemonic_forwards_matches _ = ( False )\ definition mul_mnemonic_backwards_matches :: \ string \ bool \ where \ mul_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''mul'')))) then True else if (((p00 = (''mulh'')))) then True else if (((p00 = (''mulhsu'')))) then True else if (((p00 = (''mulhu'')))) then True else False))\ for arg1 :: " string " \ \\val _s1046_ : string -> maybe string\\ definition s1046 :: \ string \(string)option \ where \ s1046 s10470 = ( (let s10480 = s10470 in if ((string_startswith s10480 (''mulhu''))) then (case ((string_drop s10480 ((string_length (''mulhu''))))) of s1 => Some s1 ) else None))\ for s10470 :: " string " \ \\val _s1042_ : string -> maybe string\\ definition s1042 :: \ string \(string)option \ where \ s1042 s10430 = ( (let s10440 = s10430 in if ((string_startswith s10440 (''mulhsu''))) then (case ((string_drop s10440 ((string_length (''mulhsu''))))) of s1 => Some s1 ) else None))\ for s10430 :: " string " \ \\val _s1038_ : string -> maybe string\\ definition s1038 :: \ string \(string)option \ where \ s1038 s10390 = ( (let s10400 = s10390 in if ((string_startswith s10400 (''mulh''))) then (case ((string_drop s10400 ((string_length (''mulh''))))) of s1 => Some s1 ) else None))\ for s10390 :: " string " \ \\val _s1034_ : string -> maybe string\\ definition s1034 :: \ string \(string)option \ where \ s1034 s10350 = ( (let s10360 = s10350 in if ((string_startswith s10360 (''mul''))) then (case ((string_drop s10360 ((string_length (''mul''))))) of s1 => Some s1 ) else None))\ for s10350 :: " string " definition mul_mnemonic_matches_prefix :: \ string \((bool*bool*bool)*int)option \ where \ mul_mnemonic_matches_prefix arg1 = ( (let s10370 = arg1 in if ((case ((s1034 s10370)) of Some (s1) => True | _ => False )) then (case s1034 s10370 of (Some (s1)) => Some ((False, True, True), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1038 s10370)) of Some (s1) => True | _ => False )) then (case s1038 s10370 of (Some (s1)) => Some ((True, True, True), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1042 s10370)) of Some (s1) => True | _ => False )) then (case s1042 s10370 of (Some (s1)) => Some ((True, True, False), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1046 s10370)) of Some (s1) => True | _ => False )) then (case s1046 s10370 of (Some (s1)) => Some ((True, False, False), ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val maybe_not_u_forwards : bool -> string\\ \ \\val maybe_not_u_backwards : string -> M bool\\ \ \\val maybe_not_u_forwards_matches : bool -> bool\\ \ \\val maybe_not_u_backwards_matches : string -> bool\\ \ \\val maybe_not_u_matches_prefix : string -> maybe ((bool * ii))\\ fun maybe_not_u_forwards :: \ bool \ string \ where \ maybe_not_u_forwards False = ( (''u''))\ |\ maybe_not_u_forwards True = ( (''''))\ definition maybe_not_u_backwards :: \ string \((register_value),(bool),(exception))monad \ where \ maybe_not_u_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''u'')))) then return False else if (((p00 = ('''')))) then return True else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun maybe_not_u_forwards_matches :: \ bool \ bool \ where \ maybe_not_u_forwards_matches False = ( True )\ |\ maybe_not_u_forwards_matches True = ( True )\ definition maybe_not_u_backwards_matches :: \ string \ bool \ where \ maybe_not_u_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''u'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s1054_ : string -> maybe string\\ definition s1054 :: \ string \(string)option \ where \ s1054 s10550 = ( (let s10560 = s10550 in if ((string_startswith s10560 (''''))) then (case ((string_drop s10560 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s10550 :: " string " \ \\val _s1050_ : string -> maybe string\\ definition s1050 :: \ string \(string)option \ where \ s1050 s10510 = ( (let s10520 = s10510 in if ((string_startswith s10520 (''u''))) then (case ((string_drop s10520 ((string_length (''u''))))) of s1 => Some s1 ) else None))\ for s10510 :: " string " definition maybe_not_u_matches_prefix :: \ string \(bool*int)option \ where \ maybe_not_u_matches_prefix arg1 = ( (let s10530 = arg1 in if ((case ((s1050 s10530)) of Some (s1) => True | _ => False )) then (case s1050 s10530 of (Some (s1)) => Some (False, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1054 s10530)) of Some (s1) => True | _ => False )) then (case s1054 s10530 of (Some (s1)) => Some (True, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val encdec_csrop_forwards : csrop -> mword ty2\\ \ \\val encdec_csrop_backwards : mword ty2 -> M csrop\\ \ \\val encdec_csrop_forwards_matches : csrop -> bool\\ \ \\val encdec_csrop_backwards_matches : mword ty2 -> bool\\ fun encdec_csrop_forwards :: \ csrop \(2)Word.word \ where \ encdec_csrop_forwards CSRRW = ( ( 0b01 :: 2 Word.word))\ |\ encdec_csrop_forwards CSRRS = ( ( 0b10 :: 2 Word.word))\ |\ encdec_csrop_forwards CSRRC = ( ( 0b11 :: 2 Word.word))\ definition encdec_csrop_backwards :: \(2)Word.word \((register_value),(csrop),(exception))monad \ where \ encdec_csrop_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b01 :: 2 Word.word)))) then return CSRRW else if (((b__0 = ( 0b10 :: 2 Word.word)))) then return CSRRS else if (((b__0 = ( 0b11 :: 2 Word.word)))) then return CSRRC else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(2)Word.word " fun encdec_csrop_forwards_matches :: \ csrop \ bool \ where \ encdec_csrop_forwards_matches CSRRW = ( True )\ |\ encdec_csrop_forwards_matches CSRRS = ( True )\ |\ encdec_csrop_forwards_matches CSRRC = ( True )\ definition encdec_csrop_backwards_matches :: \(2)Word.word \ bool \ where \ encdec_csrop_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b01 :: 2 Word.word)))) then True else if (((b__0 = ( 0b10 :: 2 Word.word)))) then True else if (((b__0 = ( 0b11 :: 2 Word.word)))) then True else False))\ for arg1 :: "(2)Word.word " \ \\val readCSR : mword ty12 -> M (mword ty32)\\ definition readCSR :: \(12)Word.word \((register_value),((32)Word.word),(exception))monad \ where \ readCSR csr = ( (case (csr, (( 32 :: int)::ii)) of (b__0, g__325) => if (((b__0 = ( 0xF11 :: 12 Word.word)))) then (read_reg mvendorid_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((EXTZ (( 32 :: int)::ii) w__0 :: 32 Word.word)))) else if (((b__0 = ( 0xF12 :: 12 Word.word)))) then (read_reg marchid_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0xF13 :: 12 Word.word)))) then (read_reg mimpid_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0xF14 :: 12 Word.word)))) then (read_reg mhartid_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x300 :: 12 Word.word)))) then read_reg mstatus_ref \ ((\ (w__4 :: Mstatus) . return ((get_Mstatus_bits w__4 :: 32 Word.word)))) else if (((b__0 = ( 0x301 :: 12 Word.word)))) then read_reg misa_ref \ ((\ (w__5 :: Misa) . return ((get_Misa_bits w__5 :: 32 Word.word)))) else if (((b__0 = ( 0x302 :: 12 Word.word)))) then read_reg medeleg_ref \ ((\ (w__6 :: Medeleg) . return ((get_Medeleg_bits w__6 :: 32 Word.word)))) else if (((b__0 = ( 0x303 :: 12 Word.word)))) then read_reg mideleg_ref \ ((\ (w__7 :: Minterrupts) . return ((get_Minterrupts_bits w__7 :: 32 Word.word)))) else if (((b__0 = ( 0x304 :: 12 Word.word)))) then read_reg mie_ref \ ((\ (w__8 :: Minterrupts) . return ((get_Minterrupts_bits w__8 :: 32 Word.word)))) else if (((b__0 = ( 0x305 :: 12 Word.word)))) then (get_mtvec () :: ( 32 Word.word) M) else if (((b__0 = ( 0x306 :: 12 Word.word)))) then read_reg mcounteren_ref \ ((\ (w__10 :: Counteren) . return ((EXTZ (( 32 :: int)::ii) ((get_Counteren_bits w__10 :: 32 Word.word)) :: 32 Word.word)))) else if ((((((b__0 = ( 0x310 :: 12 Word.word)))) \ (((g__325 = (( 32 :: int)::ii))))))) then read_reg mstatush_ref \ ((\ (w__11 :: Mstatush) . return ((get_Mstatush_bits w__11 :: 32 Word.word)))) else if (((b__0 = ( 0x320 :: 12 Word.word)))) then read_reg mcountinhibit_ref \ ((\ (w__12 :: Counterin) . return ((EXTZ (( 32 :: int)::ii) ((get_Counterin_bits w__12 :: 32 Word.word)) :: 32 Word.word)))) else if (((b__0 = ( 0x340 :: 12 Word.word)))) then (read_reg mscratch_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x341 :: 12 Word.word)))) then (get_xret_target Machine :: ( 32 Word.word) M) \ ((\ (w__14 :: 32 Word.word) . (pc_alignment_mask () :: ( 32 Word.word) M) \ ((\ (w__15 :: 32 Word.word) . return ((and_vec w__14 w__15 :: 32 Word.word)))))) else if (((b__0 = ( 0x342 :: 12 Word.word)))) then read_reg mcause_ref \ ((\ (w__16 :: Mcause) . return ((get_Mcause_bits w__16 :: 32 Word.word)))) else if (((b__0 = ( 0x343 :: 12 Word.word)))) then (read_reg mtval_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x344 :: 12 Word.word)))) then read_reg mip_ref \ ((\ (w__18 :: Minterrupts) . return ((get_Minterrupts_bits w__18 :: 32 Word.word)))) else if (((b__0 = ( 0x3A0 :: 12 Word.word)))) then (pmpReadCfgReg (( 0 :: int)::ii) :: ( 32 Word.word) M) else if ((((((b__0 = ( 0x3A1 :: 12 Word.word)))) \ (((g__325 = (( 32 :: int)::ii))))))) then (pmpReadCfgReg (( 1 :: int)::ii) :: ( 32 Word.word) M) else if (((b__0 = ( 0x3A2 :: 12 Word.word)))) then (pmpReadCfgReg (( 2 :: int)::ii) :: ( 32 Word.word) M) else if ((((((b__0 = ( 0x3A3 :: 12 Word.word)))) \ (((g__325 = (( 32 :: int)::ii))))))) then (pmpReadCfgReg (( 3 :: int)::ii) :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B0 :: 12 Word.word)))) then (read_reg pmpaddr0_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B1 :: 12 Word.word)))) then (read_reg pmpaddr1_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B2 :: 12 Word.word)))) then (read_reg pmpaddr2_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B3 :: 12 Word.word)))) then (read_reg pmpaddr3_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B4 :: 12 Word.word)))) then (read_reg pmpaddr4_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B5 :: 12 Word.word)))) then (read_reg pmpaddr5_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B6 :: 12 Word.word)))) then (read_reg pmpaddr6_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B7 :: 12 Word.word)))) then (read_reg pmpaddr7_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B8 :: 12 Word.word)))) then (read_reg pmpaddr8_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3B9 :: 12 Word.word)))) then (read_reg pmpaddr9_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3BA :: 12 Word.word)))) then (read_reg pmpaddr10_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3BB :: 12 Word.word)))) then (read_reg pmpaddr11_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3BC :: 12 Word.word)))) then (read_reg pmpaddr12_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3BD :: 12 Word.word)))) then (read_reg pmpaddr13_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3BE :: 12 Word.word)))) then (read_reg pmpaddr14_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x3BF :: 12 Word.word)))) then (read_reg pmpaddr15_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0xB00 :: 12 Word.word)))) then (read_reg mcycle_ref :: ( 64 Word.word) M) \ ((\ (w__39 :: 64 Word.word) . return ((subrange_vec_dec w__39 (((( 32 :: int)::ii) - (( 1 :: int)::ii))) (( 0 :: int)::ii) :: 32 Word.word)))) else if (((b__0 = ( 0xB02 :: 12 Word.word)))) then (read_reg minstret_ref :: ( 64 Word.word) M) \ ((\ (w__40 :: 64 Word.word) . return ((subrange_vec_dec w__40 (((( 32 :: int)::ii) - (( 1 :: int)::ii))) (( 0 :: int)::ii) :: 32 Word.word)))) else if ((((((b__0 = ( 0xB80 :: 12 Word.word)))) \ (((g__325 = (( 32 :: int)::ii))))))) then (read_reg mcycle_ref :: ( 64 Word.word) M) \ ((\ (w__41 :: 64 Word.word) . return ((subrange_vec_dec w__41 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)))) else if ((((((b__0 = ( 0xB82 :: 12 Word.word)))) \ (((g__325 = (( 32 :: int)::ii))))))) then (read_reg minstret_ref :: ( 64 Word.word) M) \ ((\ (w__42 :: 64 Word.word) . return ((subrange_vec_dec w__42 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)))) else if (((b__0 = ( 0x7A0 :: 12 Word.word)))) then (read_reg tselect_ref :: ( 32 Word.word) M) \ ((\ (w__43 :: 32 Word.word) . return ((not_vec w__43 :: 32 Word.word)))) else if (((b__0 = ( 0x100 :: 12 Word.word)))) then read_reg mstatus_ref \ ((\ (w__44 :: Mstatus) . return ((get_Sstatus_bits ((lower_mstatus w__44)) :: 32 Word.word)))) else if (((b__0 = ( 0x102 :: 12 Word.word)))) then read_reg sedeleg_ref \ ((\ (w__45 :: Sedeleg) . return ((get_Sedeleg_bits w__45 :: 32 Word.word)))) else if (((b__0 = ( 0x103 :: 12 Word.word)))) then read_reg sideleg_ref \ ((\ (w__46 :: Sinterrupts) . return ((get_Sinterrupts_bits w__46 :: 32 Word.word)))) else if (((b__0 = ( 0x104 :: 12 Word.word)))) then read_reg mie_ref \ ((\ (w__47 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__48 :: Minterrupts) . return ((get_Sinterrupts_bits ((lower_mie w__47 w__48)) :: 32 Word.word)))))) else if (((b__0 = ( 0x105 :: 12 Word.word)))) then (get_stvec () :: ( 32 Word.word) M) else if (((b__0 = ( 0x106 :: 12 Word.word)))) then read_reg scounteren_ref \ ((\ (w__50 :: Counteren) . return ((EXTZ (( 32 :: int)::ii) ((get_Counteren_bits w__50 :: 32 Word.word)) :: 32 Word.word)))) else if (((b__0 = ( 0x140 :: 12 Word.word)))) then (read_reg sscratch_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x141 :: 12 Word.word)))) then (get_xret_target Supervisor :: ( 32 Word.word) M) \ ((\ (w__52 :: 32 Word.word) . (pc_alignment_mask () :: ( 32 Word.word) M) \ ((\ (w__53 :: 32 Word.word) . return ((and_vec w__52 w__53 :: 32 Word.word)))))) else if (((b__0 = ( 0x142 :: 12 Word.word)))) then read_reg scause_ref \ ((\ (w__54 :: Mcause) . return ((get_Mcause_bits w__54 :: 32 Word.word)))) else if (((b__0 = ( 0x143 :: 12 Word.word)))) then (read_reg stval_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0x144 :: 12 Word.word)))) then read_reg mip_ref \ ((\ (w__56 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__57 :: Minterrupts) . return ((get_Sinterrupts_bits ((lower_mip w__56 w__57)) :: 32 Word.word)))))) else if (((b__0 = ( 0x180 :: 12 Word.word)))) then (read_reg satp_ref :: ( 32 Word.word) M) else if (((b__0 = ( 0xC00 :: 12 Word.word)))) then (read_reg mcycle_ref :: ( 64 Word.word) M) \ ((\ (w__59 :: 64 Word.word) . return ((subrange_vec_dec w__59 (((( 32 :: int)::ii) - (( 1 :: int)::ii))) (( 0 :: int)::ii) :: 32 Word.word)))) else if (((b__0 = ( 0xC01 :: 12 Word.word)))) then (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__60 :: 64 Word.word) . return ((subrange_vec_dec w__60 (((( 32 :: int)::ii) - (( 1 :: int)::ii))) (( 0 :: int)::ii) :: 32 Word.word)))) else if (((b__0 = ( 0xC02 :: 12 Word.word)))) then (read_reg minstret_ref :: ( 64 Word.word) M) \ ((\ (w__61 :: 64 Word.word) . return ((subrange_vec_dec w__61 (((( 32 :: int)::ii) - (( 1 :: int)::ii))) (( 0 :: int)::ii) :: 32 Word.word)))) else if ((((((b__0 = ( 0xC80 :: 12 Word.word)))) \ (((g__325 = (( 32 :: int)::ii))))))) then (read_reg mcycle_ref :: ( 64 Word.word) M) \ ((\ (w__62 :: 64 Word.word) . return ((subrange_vec_dec w__62 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)))) else if ((((((b__0 = ( 0xC81 :: 12 Word.word)))) \ (((g__325 = (( 32 :: int)::ii))))))) then (read_reg mtime_ref :: ( 64 Word.word) M) \ ((\ (w__63 :: 64 Word.word) . return ((subrange_vec_dec w__63 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)))) else if ((((((b__0 = ( 0xC82 :: 12 Word.word)))) \ (((g__325 = (( 32 :: int)::ii))))))) then (read_reg minstret_ref :: ( 64 Word.word) M) \ ((\ (w__64 :: 64 Word.word) . return ((subrange_vec_dec w__64 (( 63 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)))) else (ext_read_CSR csr :: ( ( 32 Word.word)option) M) \ ((\ (w__65 :: ( 32 Word.word)option) . return ((case w__65 of Some (res) => res | None => (let (_ :: unit) = (print_bits0 (''unhandled read to CSR '') csr) in (EXTZ (( 32 :: int)::ii) ( 0x0 :: 4 Word.word) :: 32 Word.word)) )))) ) \ ((\ (res :: xlenbits) . (let (_ :: unit) = (if ((get_config_print_reg () )) then print_dbg (((@) (''CSR '') (((@) ((csr_name csr)) (((@) ('' -> '') ((string_of_bits res)))))))) else () ) in return res))))\ for csr :: "(12)Word.word " \ \\val writeCSR : mword ty12 -> mword ty32 -> M unit\\ definition writeCSR :: \(12)Word.word \(32)Word.word \((register_value),(unit),(exception))monad \ where \ writeCSR (csr :: csreg) (value1 :: xlenbits) = ( (case (csr, (( 32 :: int)::ii)) of (b__0, g__324) => if (((b__0 = ( 0x300 :: 12 Word.word)))) then read_reg mstatus_ref \ ((\ (w__0 :: Mstatus) . legalize_mstatus w__0 value1 \ ((\ (w__1 :: Mstatus) . (write_reg mstatus_ref w__1 \ read_reg mstatus_ref) \ ((\ (w__2 :: Mstatus) . return (Some ((get_Mstatus_bits w__2 :: 32 Word.word))))))))) else if (((b__0 = ( 0x301 :: 12 Word.word)))) then read_reg misa_ref \ ((\ (w__3 :: Misa) . legalize_misa w__3 value1 \ ((\ (w__4 :: Misa) . (write_reg misa_ref w__4 \ read_reg misa_ref) \ ((\ (w__5 :: Misa) . return (Some ((get_Misa_bits w__5 :: 32 Word.word))))))))) else if (((b__0 = ( 0x302 :: 12 Word.word)))) then read_reg medeleg_ref \ ((\ (w__6 :: Medeleg) . (write_reg medeleg_ref ((legalize_medeleg w__6 value1)) \ read_reg medeleg_ref) \ ((\ (w__7 :: Medeleg) . return (Some ((get_Medeleg_bits w__7 :: 32 Word.word))))))) else if (((b__0 = ( 0x303 :: 12 Word.word)))) then read_reg mideleg_ref \ ((\ (w__8 :: Minterrupts) . (write_reg mideleg_ref ((legalize_mideleg w__8 value1)) \ read_reg mideleg_ref) \ ((\ (w__9 :: Minterrupts) . return (Some ((get_Minterrupts_bits w__9 :: 32 Word.word))))))) else if (((b__0 = ( 0x304 :: 12 Word.word)))) then read_reg mie_ref \ ((\ (w__10 :: Minterrupts) . legalize_mie w__10 value1 \ ((\ (w__11 :: Minterrupts) . (write_reg mie_ref w__11 \ read_reg mie_ref) \ ((\ (w__12 :: Minterrupts) . return (Some ((get_Minterrupts_bits w__12 :: 32 Word.word))))))))) else if (((b__0 = ( 0x305 :: 12 Word.word)))) then (set_mtvec value1 :: ( 32 Word.word) M) \ ((\ (w__13 :: 32 Word.word) . return (Some w__13))) else if (((b__0 = ( 0x306 :: 12 Word.word)))) then read_reg mcounteren_ref \ ((\ (w__14 :: Counteren) . (write_reg mcounteren_ref ((legalize_mcounteren w__14 value1)) \ read_reg mcounteren_ref) \ ((\ (w__15 :: Counteren) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Counteren_bits w__15 :: 32 Word.word)) :: 32 Word.word))))))) else if ((((((b__0 = ( 0x310 :: 12 Word.word)))) \ (((g__324 = (( 32 :: int)::ii))))))) then read_reg mstatush_ref \ ((\ (w__16 :: Mstatush) . return (Some ((get_Mstatush_bits w__16 :: 32 Word.word))))) else if (((b__0 = ( 0x320 :: 12 Word.word)))) then read_reg mcountinhibit_ref \ ((\ (w__17 :: Counterin) . (write_reg mcountinhibit_ref ((legalize_mcountinhibit w__17 value1)) \ read_reg mcountinhibit_ref) \ ((\ (w__18 :: Counterin) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Counterin_bits w__18 :: 32 Word.word)) :: 32 Word.word))))))) else if (((b__0 = ( 0x340 :: 12 Word.word)))) then (write_reg mscratch_ref value1 \ (read_reg mscratch_ref :: ( 32 Word.word) M)) \ ((\ (w__19 :: 32 Word.word) . return (Some w__19))) else if (((b__0 = ( 0x341 :: 12 Word.word)))) then (set_xret_target Machine value1 :: ( 32 Word.word) M) \ ((\ (w__20 :: 32 Word.word) . return (Some w__20))) else if (((b__0 = ( 0x342 :: 12 Word.word)))) then (set_Mcause_bits mcause_ref value1 \ read_reg mcause_ref) \ ((\ (w__21 :: Mcause) . return (Some ((get_Mcause_bits w__21 :: 32 Word.word))))) else if (((b__0 = ( 0x343 :: 12 Word.word)))) then (write_reg mtval_ref value1 \ (read_reg mtval_ref :: ( 32 Word.word) M)) \ ((\ (w__22 :: 32 Word.word) . return (Some w__22))) else if (((b__0 = ( 0x344 :: 12 Word.word)))) then read_reg mip_ref \ ((\ (w__23 :: Minterrupts) . legalize_mip w__23 value1 \ ((\ (w__24 :: Minterrupts) . (write_reg mip_ref w__24 \ read_reg mip_ref) \ ((\ (w__25 :: Minterrupts) . return (Some ((get_Minterrupts_bits w__25 :: 32 Word.word))))))))) else if (((b__0 = ( 0x3A0 :: 12 Word.word)))) then (pmpWriteCfgReg (( 0 :: int)::ii) value1 \ (pmpReadCfgReg (( 0 :: int)::ii) :: ( 32 Word.word) M)) \ ((\ (w__26 :: 32 Word.word) . return (Some w__26))) else if ((((((b__0 = ( 0x3A1 :: 12 Word.word)))) \ (((g__324 = (( 32 :: int)::ii))))))) then (pmpWriteCfgReg (( 1 :: int)::ii) value1 \ (pmpReadCfgReg (( 1 :: int)::ii) :: ( 32 Word.word) M)) \ ((\ (w__27 :: 32 Word.word) . return (Some w__27))) else if (((b__0 = ( 0x3A2 :: 12 Word.word)))) then (pmpWriteCfgReg (( 2 :: int)::ii) value1 \ (pmpReadCfgReg (( 2 :: int)::ii) :: ( 32 Word.word) M)) \ ((\ (w__28 :: 32 Word.word) . return (Some w__28))) else if ((((((b__0 = ( 0x3A3 :: 12 Word.word)))) \ (((g__324 = (( 32 :: int)::ii))))))) then (pmpWriteCfgReg (( 3 :: int)::ii) value1 \ (pmpReadCfgReg (( 3 :: int)::ii) :: ( 32 Word.word) M)) \ ((\ (w__29 :: 32 Word.word) . return (Some w__29))) else if (((b__0 = ( 0x3B0 :: 12 Word.word)))) then read_reg pmp0cfg_ref \ ((\ (w__30 :: Pmpcfg_ent) . read_reg pmp1cfg_ref \ ((\ (w__31 :: Pmpcfg_ent) . pmpTORLocked w__31 \ ((\ (w__32 :: bool) . (read_reg pmpaddr0_ref :: ( 32 Word.word) M) \ ((\ (w__33 :: 32 Word.word) . (write_reg pmpaddr0_ref ((pmpWriteAddr ((pmpLocked w__30)) w__32 w__33 value1 :: 32 Word.word)) \ (read_reg pmpaddr0_ref :: ( 32 Word.word) M)) \ ((\ (w__34 :: 32 Word.word) . return (Some w__34))))))))))) else if (((b__0 = ( 0x3B1 :: 12 Word.word)))) then read_reg pmp1cfg_ref \ ((\ (w__35 :: Pmpcfg_ent) . read_reg pmp2cfg_ref \ ((\ (w__36 :: Pmpcfg_ent) . pmpTORLocked w__36 \ ((\ (w__37 :: bool) . (read_reg pmpaddr1_ref :: ( 32 Word.word) M) \ ((\ (w__38 :: 32 Word.word) . (write_reg pmpaddr1_ref ((pmpWriteAddr ((pmpLocked w__35)) w__37 w__38 value1 :: 32 Word.word)) \ (read_reg pmpaddr1_ref :: ( 32 Word.word) M)) \ ((\ (w__39 :: 32 Word.word) . return (Some w__39))))))))))) else if (((b__0 = ( 0x3B2 :: 12 Word.word)))) then read_reg pmp2cfg_ref \ ((\ (w__40 :: Pmpcfg_ent) . read_reg pmp3cfg_ref \ ((\ (w__41 :: Pmpcfg_ent) . pmpTORLocked w__41 \ ((\ (w__42 :: bool) . (read_reg pmpaddr2_ref :: ( 32 Word.word) M) \ ((\ (w__43 :: 32 Word.word) . (write_reg pmpaddr2_ref ((pmpWriteAddr ((pmpLocked w__40)) w__42 w__43 value1 :: 32 Word.word)) \ (read_reg pmpaddr2_ref :: ( 32 Word.word) M)) \ ((\ (w__44 :: 32 Word.word) . return (Some w__44))))))))))) else if (((b__0 = ( 0x3B3 :: 12 Word.word)))) then read_reg pmp3cfg_ref \ ((\ (w__45 :: Pmpcfg_ent) . read_reg pmp4cfg_ref \ ((\ (w__46 :: Pmpcfg_ent) . pmpTORLocked w__46 \ ((\ (w__47 :: bool) . (read_reg pmpaddr3_ref :: ( 32 Word.word) M) \ ((\ (w__48 :: 32 Word.word) . (write_reg pmpaddr3_ref ((pmpWriteAddr ((pmpLocked w__45)) w__47 w__48 value1 :: 32 Word.word)) \ (read_reg pmpaddr3_ref :: ( 32 Word.word) M)) \ ((\ (w__49 :: 32 Word.word) . return (Some w__49))))))))))) else if (((b__0 = ( 0x3B4 :: 12 Word.word)))) then read_reg pmp4cfg_ref \ ((\ (w__50 :: Pmpcfg_ent) . read_reg pmp5cfg_ref \ ((\ (w__51 :: Pmpcfg_ent) . pmpTORLocked w__51 \ ((\ (w__52 :: bool) . (read_reg pmpaddr4_ref :: ( 32 Word.word) M) \ ((\ (w__53 :: 32 Word.word) . (write_reg pmpaddr4_ref ((pmpWriteAddr ((pmpLocked w__50)) w__52 w__53 value1 :: 32 Word.word)) \ (read_reg pmpaddr4_ref :: ( 32 Word.word) M)) \ ((\ (w__54 :: 32 Word.word) . return (Some w__54))))))))))) else if (((b__0 = ( 0x3B5 :: 12 Word.word)))) then read_reg pmp5cfg_ref \ ((\ (w__55 :: Pmpcfg_ent) . read_reg pmp6cfg_ref \ ((\ (w__56 :: Pmpcfg_ent) . pmpTORLocked w__56 \ ((\ (w__57 :: bool) . (read_reg pmpaddr5_ref :: ( 32 Word.word) M) \ ((\ (w__58 :: 32 Word.word) . (write_reg pmpaddr5_ref ((pmpWriteAddr ((pmpLocked w__55)) w__57 w__58 value1 :: 32 Word.word)) \ (read_reg pmpaddr5_ref :: ( 32 Word.word) M)) \ ((\ (w__59 :: 32 Word.word) . return (Some w__59))))))))))) else if (((b__0 = ( 0x3B6 :: 12 Word.word)))) then read_reg pmp6cfg_ref \ ((\ (w__60 :: Pmpcfg_ent) . read_reg pmp7cfg_ref \ ((\ (w__61 :: Pmpcfg_ent) . pmpTORLocked w__61 \ ((\ (w__62 :: bool) . (read_reg pmpaddr6_ref :: ( 32 Word.word) M) \ ((\ (w__63 :: 32 Word.word) . (write_reg pmpaddr6_ref ((pmpWriteAddr ((pmpLocked w__60)) w__62 w__63 value1 :: 32 Word.word)) \ (read_reg pmpaddr6_ref :: ( 32 Word.word) M)) \ ((\ (w__64 :: 32 Word.word) . return (Some w__64))))))))))) else if (((b__0 = ( 0x3B7 :: 12 Word.word)))) then read_reg pmp7cfg_ref \ ((\ (w__65 :: Pmpcfg_ent) . read_reg pmp8cfg_ref \ ((\ (w__66 :: Pmpcfg_ent) . pmpTORLocked w__66 \ ((\ (w__67 :: bool) . (read_reg pmpaddr7_ref :: ( 32 Word.word) M) \ ((\ (w__68 :: 32 Word.word) . (write_reg pmpaddr7_ref ((pmpWriteAddr ((pmpLocked w__65)) w__67 w__68 value1 :: 32 Word.word)) \ (read_reg pmpaddr7_ref :: ( 32 Word.word) M)) \ ((\ (w__69 :: 32 Word.word) . return (Some w__69))))))))))) else if (((b__0 = ( 0x3B8 :: 12 Word.word)))) then read_reg pmp8cfg_ref \ ((\ (w__70 :: Pmpcfg_ent) . read_reg pmp9cfg_ref \ ((\ (w__71 :: Pmpcfg_ent) . pmpTORLocked w__71 \ ((\ (w__72 :: bool) . (read_reg pmpaddr8_ref :: ( 32 Word.word) M) \ ((\ (w__73 :: 32 Word.word) . (write_reg pmpaddr8_ref ((pmpWriteAddr ((pmpLocked w__70)) w__72 w__73 value1 :: 32 Word.word)) \ (read_reg pmpaddr8_ref :: ( 32 Word.word) M)) \ ((\ (w__74 :: 32 Word.word) . return (Some w__74))))))))))) else if (((b__0 = ( 0x3B9 :: 12 Word.word)))) then read_reg pmp9cfg_ref \ ((\ (w__75 :: Pmpcfg_ent) . read_reg pmp10cfg_ref \ ((\ (w__76 :: Pmpcfg_ent) . pmpTORLocked w__76 \ ((\ (w__77 :: bool) . (read_reg pmpaddr9_ref :: ( 32 Word.word) M) \ ((\ (w__78 :: 32 Word.word) . (write_reg pmpaddr9_ref ((pmpWriteAddr ((pmpLocked w__75)) w__77 w__78 value1 :: 32 Word.word)) \ (read_reg pmpaddr9_ref :: ( 32 Word.word) M)) \ ((\ (w__79 :: 32 Word.word) . return (Some w__79))))))))))) else if (((b__0 = ( 0x3BA :: 12 Word.word)))) then read_reg pmp10cfg_ref \ ((\ (w__80 :: Pmpcfg_ent) . read_reg pmp11cfg_ref \ ((\ (w__81 :: Pmpcfg_ent) . pmpTORLocked w__81 \ ((\ (w__82 :: bool) . (read_reg pmpaddr10_ref :: ( 32 Word.word) M) \ ((\ (w__83 :: 32 Word.word) . (write_reg pmpaddr10_ref ((pmpWriteAddr ((pmpLocked w__80)) w__82 w__83 value1 :: 32 Word.word)) \ (read_reg pmpaddr10_ref :: ( 32 Word.word) M)) \ ((\ (w__84 :: 32 Word.word) . return (Some w__84))))))))))) else if (((b__0 = ( 0x3BB :: 12 Word.word)))) then read_reg pmp11cfg_ref \ ((\ (w__85 :: Pmpcfg_ent) . read_reg pmp12cfg_ref \ ((\ (w__86 :: Pmpcfg_ent) . pmpTORLocked w__86 \ ((\ (w__87 :: bool) . (read_reg pmpaddr11_ref :: ( 32 Word.word) M) \ ((\ (w__88 :: 32 Word.word) . (write_reg pmpaddr11_ref ((pmpWriteAddr ((pmpLocked w__85)) w__87 w__88 value1 :: 32 Word.word)) \ (read_reg pmpaddr11_ref :: ( 32 Word.word) M)) \ ((\ (w__89 :: 32 Word.word) . return (Some w__89))))))))))) else if (((b__0 = ( 0x3BC :: 12 Word.word)))) then read_reg pmp12cfg_ref \ ((\ (w__90 :: Pmpcfg_ent) . read_reg pmp13cfg_ref \ ((\ (w__91 :: Pmpcfg_ent) . pmpTORLocked w__91 \ ((\ (w__92 :: bool) . (read_reg pmpaddr12_ref :: ( 32 Word.word) M) \ ((\ (w__93 :: 32 Word.word) . (write_reg pmpaddr12_ref ((pmpWriteAddr ((pmpLocked w__90)) w__92 w__93 value1 :: 32 Word.word)) \ (read_reg pmpaddr12_ref :: ( 32 Word.word) M)) \ ((\ (w__94 :: 32 Word.word) . return (Some w__94))))))))))) else if (((b__0 = ( 0x3BD :: 12 Word.word)))) then read_reg pmp13cfg_ref \ ((\ (w__95 :: Pmpcfg_ent) . read_reg pmp14cfg_ref \ ((\ (w__96 :: Pmpcfg_ent) . pmpTORLocked w__96 \ ((\ (w__97 :: bool) . (read_reg pmpaddr13_ref :: ( 32 Word.word) M) \ ((\ (w__98 :: 32 Word.word) . (write_reg pmpaddr13_ref ((pmpWriteAddr ((pmpLocked w__95)) w__97 w__98 value1 :: 32 Word.word)) \ (read_reg pmpaddr13_ref :: ( 32 Word.word) M)) \ ((\ (w__99 :: 32 Word.word) . return (Some w__99))))))))))) else if (((b__0 = ( 0x3BE :: 12 Word.word)))) then read_reg pmp14cfg_ref \ ((\ (w__100 :: Pmpcfg_ent) . read_reg pmp15cfg_ref \ ((\ (w__101 :: Pmpcfg_ent) . pmpTORLocked w__101 \ ((\ (w__102 :: bool) . (read_reg pmpaddr14_ref :: ( 32 Word.word) M) \ ((\ (w__103 :: 32 Word.word) . (write_reg pmpaddr14_ref ((pmpWriteAddr ((pmpLocked w__100)) w__102 w__103 value1 :: 32 Word.word)) \ (read_reg pmpaddr14_ref :: ( 32 Word.word) M)) \ ((\ (w__104 :: 32 Word.word) . return (Some w__104))))))))))) else if (((b__0 = ( 0x3BF :: 12 Word.word)))) then read_reg pmp15cfg_ref \ ((\ (w__105 :: Pmpcfg_ent) . (read_reg pmpaddr15_ref :: ( 32 Word.word) M) \ ((\ (w__106 :: 32 Word.word) . (write_reg pmpaddr15_ref ((pmpWriteAddr ((pmpLocked w__105)) False w__106 value1 :: 32 Word.word)) \ (read_reg pmpaddr15_ref :: ( 32 Word.word) M)) \ ((\ (w__107 :: 32 Word.word) . return (Some w__107))))))) else if (((b__0 = ( 0xB00 :: 12 Word.word)))) then (read_reg mcycle_ref :: ( 64 Word.word) M) \ ((\ (w__108 :: 64 Word.word) . write_reg mcycle_ref ((update_subrange_vec_dec w__108 (((( 32 :: int)::ii) - (( 1 :: int)::ii))) (( 0 :: int)::ii) value1 :: 64 Word.word)) \ return (Some value1))) else if (((b__0 = ( 0xB02 :: 12 Word.word)))) then (read_reg minstret_ref :: ( 64 Word.word) M) \ ((\ (w__109 :: 64 Word.word) . (write_reg minstret_ref ((update_subrange_vec_dec w__109 (((( 32 :: int)::ii) - (( 1 :: int)::ii))) (( 0 :: int)::ii) value1 :: 64 Word.word)) \ write_reg minstret_written_ref True) \ return (Some value1))) else if ((((((b__0 = ( 0xB80 :: 12 Word.word)))) \ (((g__324 = (( 32 :: int)::ii))))))) then (read_reg mcycle_ref :: ( 64 Word.word) M) \ ((\ (w__110 :: 64 Word.word) . write_reg mcycle_ref ((update_subrange_vec_dec w__110 (( 63 :: int)::ii) (( 32 :: int)::ii) value1 :: 64 Word.word)) \ return (Some value1))) else if ((((((b__0 = ( 0xB82 :: 12 Word.word)))) \ (((g__324 = (( 32 :: int)::ii))))))) then (read_reg minstret_ref :: ( 64 Word.word) M) \ ((\ (w__111 :: 64 Word.word) . (write_reg minstret_ref ((update_subrange_vec_dec w__111 (( 63 :: int)::ii) (( 32 :: int)::ii) value1 :: 64 Word.word)) \ write_reg minstret_written_ref True) \ return (Some value1))) else if (((b__0 = ( 0x7A0 :: 12 Word.word)))) then (write_reg tselect_ref value1 \ (read_reg tselect_ref :: ( 32 Word.word) M)) \ ((\ (w__112 :: 32 Word.word) . return (Some w__112))) else if (((b__0 = ( 0x100 :: 12 Word.word)))) then read_reg mstatus_ref \ ((\ (w__113 :: Mstatus) . legalize_sstatus w__113 value1 \ ((\ (w__114 :: Mstatus) . (write_reg mstatus_ref w__114 \ read_reg mstatus_ref) \ ((\ (w__115 :: Mstatus) . return (Some ((get_Mstatus_bits w__115 :: 32 Word.word))))))))) else if (((b__0 = ( 0x102 :: 12 Word.word)))) then read_reg sedeleg_ref \ ((\ (w__116 :: Sedeleg) . (write_reg sedeleg_ref ((legalize_sedeleg w__116 value1)) \ read_reg sedeleg_ref) \ ((\ (w__117 :: Sedeleg) . return (Some ((get_Sedeleg_bits w__117 :: 32 Word.word))))))) else if (((b__0 = ( 0x103 :: 12 Word.word)))) then (set_Sinterrupts_bits sideleg_ref value1 \ read_reg sideleg_ref) \ ((\ (w__118 :: Sinterrupts) . return (Some ((get_Sinterrupts_bits w__118 :: 32 Word.word))))) else if (((b__0 = ( 0x104 :: 12 Word.word)))) then read_reg mie_ref \ ((\ (w__119 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__120 :: Minterrupts) . legalize_sie w__119 w__120 value1 \ ((\ (w__121 :: Minterrupts) . (write_reg mie_ref w__121 \ read_reg mie_ref) \ ((\ (w__122 :: Minterrupts) . return (Some ((get_Minterrupts_bits w__122 :: 32 Word.word))))))))))) else if (((b__0 = ( 0x105 :: 12 Word.word)))) then (set_stvec value1 :: ( 32 Word.word) M) \ ((\ (w__123 :: 32 Word.word) . return (Some w__123))) else if (((b__0 = ( 0x106 :: 12 Word.word)))) then read_reg scounteren_ref \ ((\ (w__124 :: Counteren) . (write_reg scounteren_ref ((legalize_scounteren w__124 value1)) \ read_reg scounteren_ref) \ ((\ (w__125 :: Counteren) . return (Some ((EXTZ (( 32 :: int)::ii) ((get_Counteren_bits w__125 :: 32 Word.word)) :: 32 Word.word))))))) else if (((b__0 = ( 0x140 :: 12 Word.word)))) then (write_reg sscratch_ref value1 \ (read_reg sscratch_ref :: ( 32 Word.word) M)) \ ((\ (w__126 :: 32 Word.word) . return (Some w__126))) else if (((b__0 = ( 0x141 :: 12 Word.word)))) then (set_xret_target Supervisor value1 :: ( 32 Word.word) M) \ ((\ (w__127 :: 32 Word.word) . return (Some w__127))) else if (((b__0 = ( 0x142 :: 12 Word.word)))) then (set_Mcause_bits scause_ref value1 \ read_reg scause_ref) \ ((\ (w__128 :: Mcause) . return (Some ((get_Mcause_bits w__128 :: 32 Word.word))))) else if (((b__0 = ( 0x143 :: 12 Word.word)))) then (write_reg stval_ref value1 \ (read_reg stval_ref :: ( 32 Word.word) M)) \ ((\ (w__129 :: 32 Word.word) . return (Some w__129))) else if (((b__0 = ( 0x144 :: 12 Word.word)))) then read_reg mip_ref \ ((\ (w__130 :: Minterrupts) . read_reg mideleg_ref \ ((\ (w__131 :: Minterrupts) . legalize_sip w__130 w__131 value1 \ ((\ (w__132 :: Minterrupts) . (write_reg mip_ref w__132 \ read_reg mip_ref) \ ((\ (w__133 :: Minterrupts) . return (Some ((get_Minterrupts_bits w__133 :: 32 Word.word))))))))))) else if (((b__0 = ( 0x180 :: 12 Word.word)))) then cur_Architecture () \ ((\ (w__134 :: Architecture) . (read_reg satp_ref :: ( 32 Word.word) M) \ ((\ (w__135 :: 32 Word.word) . (write_reg satp_ref ((legalize_satp w__134 w__135 value1 :: 32 Word.word)) \ (read_reg satp_ref :: ( 32 Word.word) M)) \ ((\ (w__136 :: 32 Word.word) . return (Some w__136))))))) else (ext_write_CSR csr value1 :: ( ( 32 Word.word)option) M) ) \ ((\ (res :: xlenbits option) . return ((case res of Some (v) => if ((get_config_print_reg () )) then print_dbg (((@) (''CSR '') (((@) ((csr_name csr)) (((@) ('' <- '') (((@) ((string_of_bits v)) (((@) ('' (input: '') (((@) ((string_of_bits value1)) ('')''))))))))))))) else () | None => print_bits0 (''unhandled write to CSR '') csr )))))\ for csr :: "(12)Word.word " and value1 :: "(32)Word.word " \ \\val maybe_i_forwards : bool -> string\\ \ \\val maybe_i_backwards : string -> M bool\\ \ \\val maybe_i_forwards_matches : bool -> bool\\ \ \\val maybe_i_backwards_matches : string -> bool\\ \ \\val maybe_i_matches_prefix : string -> maybe ((bool * ii))\\ fun maybe_i_forwards :: \ bool \ string \ where \ maybe_i_forwards True = ( (''i''))\ |\ maybe_i_forwards False = ( (''''))\ definition maybe_i_backwards :: \ string \((register_value),(bool),(exception))monad \ where \ maybe_i_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''i'')))) then return True else if (((p00 = ('''')))) then return False else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun maybe_i_forwards_matches :: \ bool \ bool \ where \ maybe_i_forwards_matches True = ( True )\ |\ maybe_i_forwards_matches False = ( True )\ definition maybe_i_backwards_matches :: \ string \ bool \ where \ maybe_i_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''i'')))) then True else if (((p00 = ('''')))) then True else False))\ for arg1 :: " string " \ \\val _s1062_ : string -> maybe string\\ definition s1062 :: \ string \(string)option \ where \ s1062 s10630 = ( (let s10640 = s10630 in if ((string_startswith s10640 (''''))) then (case ((string_drop s10640 ((string_length (''''))))) of s1 => Some s1 ) else None))\ for s10630 :: " string " \ \\val _s1058_ : string -> maybe string\\ definition s1058 :: \ string \(string)option \ where \ s1058 s10590 = ( (let s10600 = s10590 in if ((string_startswith s10600 (''i''))) then (case ((string_drop s10600 ((string_length (''i''))))) of s1 => Some s1 ) else None))\ for s10590 :: " string " definition maybe_i_matches_prefix :: \ string \(bool*int)option \ where \ maybe_i_matches_prefix arg1 = ( (let s10610 = arg1 in if ((case ((s1058 s10610)) of Some (s1) => True | _ => False )) then (case s1058 s10610 of (Some (s1)) => Some (True, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1062 s10610)) of Some (s1) => True | _ => False )) then (case s1062 s10610 of (Some (s1)) => Some (False, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val csr_mnemonic_forwards : csrop -> string\\ \ \\val csr_mnemonic_backwards : string -> M csrop\\ \ \\val csr_mnemonic_forwards_matches : csrop -> bool\\ \ \\val csr_mnemonic_backwards_matches : string -> bool\\ \ \\val csr_mnemonic_matches_prefix : string -> maybe ((csrop * ii))\\ fun csr_mnemonic_forwards :: \ csrop \ string \ where \ csr_mnemonic_forwards CSRRW = ( (''csrrw''))\ |\ csr_mnemonic_forwards CSRRS = ( (''csrrs''))\ |\ csr_mnemonic_forwards CSRRC = ( (''csrrc''))\ definition csr_mnemonic_backwards :: \ string \((register_value),(csrop),(exception))monad \ where \ csr_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''csrrw'')))) then return CSRRW else if (((p00 = (''csrrs'')))) then return CSRRS else if (((p00 = (''csrrc'')))) then return CSRRC else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun csr_mnemonic_forwards_matches :: \ csrop \ bool \ where \ csr_mnemonic_forwards_matches CSRRW = ( True )\ |\ csr_mnemonic_forwards_matches CSRRS = ( True )\ |\ csr_mnemonic_forwards_matches CSRRC = ( True )\ definition csr_mnemonic_backwards_matches :: \ string \ bool \ where \ csr_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''csrrw'')))) then True else if (((p00 = (''csrrs'')))) then True else if (((p00 = (''csrrc'')))) then True else False))\ for arg1 :: " string " \ \\val _s1074_ : string -> maybe string\\ definition s1074 :: \ string \(string)option \ where \ s1074 s10750 = ( (let s10760 = s10750 in if ((string_startswith s10760 (''csrrc''))) then (case ((string_drop s10760 ((string_length (''csrrc''))))) of s1 => Some s1 ) else None))\ for s10750 :: " string " \ \\val _s1070_ : string -> maybe string\\ definition s1070 :: \ string \(string)option \ where \ s1070 s10710 = ( (let s10720 = s10710 in if ((string_startswith s10720 (''csrrs''))) then (case ((string_drop s10720 ((string_length (''csrrs''))))) of s1 => Some s1 ) else None))\ for s10710 :: " string " \ \\val _s1066_ : string -> maybe string\\ definition s1066 :: \ string \(string)option \ where \ s1066 s10670 = ( (let s10680 = s10670 in if ((string_startswith s10680 (''csrrw''))) then (case ((string_drop s10680 ((string_length (''csrrw''))))) of s1 => Some s1 ) else None))\ for s10670 :: " string " definition csr_mnemonic_matches_prefix :: \ string \(csrop*int)option \ where \ csr_mnemonic_matches_prefix arg1 = ( (let s10690 = arg1 in if ((case ((s1066 s10690)) of Some (s1) => True | _ => False )) then (case s1066 s10690 of (Some (s1)) => Some (CSRRW, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1070 s10690)) of Some (s1) => True | _ => False )) then (case s1070 s10690 of (Some (s1)) => Some (CSRRS, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1074 s10690)) of Some (s1) => True | _ => False )) then (case s1074 s10690 of (Some (s1)) => Some (CSRRC, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val encdec_rounding_mode_forwards : rounding_mode -> mword ty3\\ \ \\val encdec_rounding_mode_backwards : mword ty3 -> M rounding_mode\\ \ \\val encdec_rounding_mode_forwards_matches : rounding_mode -> bool\\ \ \\val encdec_rounding_mode_backwards_matches : mword ty3 -> bool\\ fun encdec_rounding_mode_forwards :: \ rounding_mode \(3)Word.word \ where \ encdec_rounding_mode_forwards RM_RNE = ( ( 0b000 :: 3 Word.word))\ |\ encdec_rounding_mode_forwards RM_RTZ = ( ( 0b001 :: 3 Word.word))\ |\ encdec_rounding_mode_forwards RM_RDN = ( ( 0b010 :: 3 Word.word))\ |\ encdec_rounding_mode_forwards RM_RUP = ( ( 0b011 :: 3 Word.word))\ |\ encdec_rounding_mode_forwards RM_RMM = ( ( 0b100 :: 3 Word.word))\ |\ encdec_rounding_mode_forwards RM_DYN = ( ( 0b111 :: 3 Word.word))\ definition encdec_rounding_mode_backwards :: \(3)Word.word \((register_value),(rounding_mode),(exception))monad \ where \ encdec_rounding_mode_backwards arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then return RM_RNE else if (((b__0 = ( 0b001 :: 3 Word.word)))) then return RM_RTZ else if (((b__0 = ( 0b010 :: 3 Word.word)))) then return RM_RDN else if (((b__0 = ( 0b011 :: 3 Word.word)))) then return RM_RUP else if (((b__0 = ( 0b100 :: 3 Word.word)))) then return RM_RMM else if (((b__0 = ( 0b111 :: 3 Word.word)))) then return RM_DYN else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: "(3)Word.word " fun encdec_rounding_mode_forwards_matches :: \ rounding_mode \ bool \ where \ encdec_rounding_mode_forwards_matches RM_RNE = ( True )\ |\ encdec_rounding_mode_forwards_matches RM_RTZ = ( True )\ |\ encdec_rounding_mode_forwards_matches RM_RDN = ( True )\ |\ encdec_rounding_mode_forwards_matches RM_RUP = ( True )\ |\ encdec_rounding_mode_forwards_matches RM_RMM = ( True )\ |\ encdec_rounding_mode_forwards_matches RM_DYN = ( True )\ definition encdec_rounding_mode_backwards_matches :: \(3)Word.word \ bool \ where \ encdec_rounding_mode_backwards_matches arg1 = ( (let b__0 = arg1 in if (((b__0 = ( 0b000 :: 3 Word.word)))) then True else if (((b__0 = ( 0b001 :: 3 Word.word)))) then True else if (((b__0 = ( 0b010 :: 3 Word.word)))) then True else if (((b__0 = ( 0b011 :: 3 Word.word)))) then True else if (((b__0 = ( 0b100 :: 3 Word.word)))) then True else if (((b__0 = ( 0b111 :: 3 Word.word)))) then True else False))\ for arg1 :: "(3)Word.word " \ \\val frm_mnemonic_forwards : rounding_mode -> string\\ \ \\val frm_mnemonic_backwards : string -> M rounding_mode\\ \ \\val frm_mnemonic_forwards_matches : rounding_mode -> bool\\ \ \\val frm_mnemonic_backwards_matches : string -> bool\\ \ \\val frm_mnemonic_matches_prefix : string -> maybe ((rounding_mode * ii))\\ fun frm_mnemonic_forwards :: \ rounding_mode \ string \ where \ frm_mnemonic_forwards RM_RNE = ( (''rne''))\ |\ frm_mnemonic_forwards RM_RTZ = ( (''rtz''))\ |\ frm_mnemonic_forwards RM_RDN = ( (''rdn''))\ |\ frm_mnemonic_forwards RM_RUP = ( (''rup''))\ |\ frm_mnemonic_forwards RM_RMM = ( (''rmm''))\ |\ frm_mnemonic_forwards RM_DYN = ( (''dyn''))\ definition frm_mnemonic_backwards :: \ string \((register_value),(rounding_mode),(exception))monad \ where \ frm_mnemonic_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''rne'')))) then return RM_RNE else if (((p00 = (''rtz'')))) then return RM_RTZ else if (((p00 = (''rdn'')))) then return RM_RDN else if (((p00 = (''rup'')))) then return RM_RUP else if (((p00 = (''rmm'')))) then return RM_RMM else if (((p00 = (''dyn'')))) then return RM_DYN else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun frm_mnemonic_forwards_matches :: \ rounding_mode \ bool \ where \ frm_mnemonic_forwards_matches RM_RNE = ( True )\ |\ frm_mnemonic_forwards_matches RM_RTZ = ( True )\ |\ frm_mnemonic_forwards_matches RM_RDN = ( True )\ |\ frm_mnemonic_forwards_matches RM_RUP = ( True )\ |\ frm_mnemonic_forwards_matches RM_RMM = ( True )\ |\ frm_mnemonic_forwards_matches RM_DYN = ( True )\ definition frm_mnemonic_backwards_matches :: \ string \ bool \ where \ frm_mnemonic_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''rne'')))) then True else if (((p00 = (''rtz'')))) then True else if (((p00 = (''rdn'')))) then True else if (((p00 = (''rup'')))) then True else if (((p00 = (''rmm'')))) then True else if (((p00 = (''dyn'')))) then True else False))\ for arg1 :: " string " \ \\val _s1098_ : string -> maybe string\\ definition s1098 :: \ string \(string)option \ where \ s1098 s10990 = ( (let s11000 = s10990 in if ((string_startswith s11000 (''dyn''))) then (case ((string_drop s11000 ((string_length (''dyn''))))) of s1 => Some s1 ) else None))\ for s10990 :: " string " \ \\val _s1094_ : string -> maybe string\\ definition s1094 :: \ string \(string)option \ where \ s1094 s10950 = ( (let s10960 = s10950 in if ((string_startswith s10960 (''rmm''))) then (case ((string_drop s10960 ((string_length (''rmm''))))) of s1 => Some s1 ) else None))\ for s10950 :: " string " \ \\val _s1090_ : string -> maybe string\\ definition s1090 :: \ string \(string)option \ where \ s1090 s10910 = ( (let s10920 = s10910 in if ((string_startswith s10920 (''rup''))) then (case ((string_drop s10920 ((string_length (''rup''))))) of s1 => Some s1 ) else None))\ for s10910 :: " string " \ \\val _s1086_ : string -> maybe string\\ definition s1086 :: \ string \(string)option \ where \ s1086 s10870 = ( (let s10880 = s10870 in if ((string_startswith s10880 (''rdn''))) then (case ((string_drop s10880 ((string_length (''rdn''))))) of s1 => Some s1 ) else None))\ for s10870 :: " string " \ \\val _s1082_ : string -> maybe string\\ definition s1082 :: \ string \(string)option \ where \ s1082 s10830 = ( (let s10840 = s10830 in if ((string_startswith s10840 (''rtz''))) then (case ((string_drop s10840 ((string_length (''rtz''))))) of s1 => Some s1 ) else None))\ for s10830 :: " string " \ \\val _s1078_ : string -> maybe string\\ definition s1078 :: \ string \(string)option \ where \ s1078 s10790 = ( (let s10800 = s10790 in if ((string_startswith s10800 (''rne''))) then (case ((string_drop s10800 ((string_length (''rne''))))) of s1 => Some s1 ) else None))\ for s10790 :: " string " definition frm_mnemonic_matches_prefix :: \ string \(rounding_mode*int)option \ where \ frm_mnemonic_matches_prefix arg1 = ( (let s10810 = arg1 in if ((case ((s1078 s10810)) of Some (s1) => True | _ => False )) then (case s1078 s10810 of (Some (s1)) => Some (RM_RNE, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1082 s10810)) of Some (s1) => True | _ => False )) then (case s1082 s10810 of (Some (s1)) => Some (RM_RTZ, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1086 s10810)) of Some (s1) => True | _ => False )) then (case s1086 s10810 of (Some (s1)) => Some (RM_RDN, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1090 s10810)) of Some (s1) => True | _ => False )) then (case s1090 s10810 of (Some (s1)) => Some (RM_RUP, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1094 s10810)) of Some (s1) => True | _ => False )) then (case s1094 s10810 of (Some (s1)) => Some (RM_RMM, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1098 s10810)) of Some (s1) => True | _ => False )) then (case s1098 s10810 of (Some (s1)) => Some (RM_DYN, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val select_instr_or_fcsr_rm : rounding_mode -> M rounding_mode\\ definition select_instr_or_fcsr_rm :: \ rounding_mode \((register_value),(rounding_mode),(exception))monad \ where \ select_instr_or_fcsr_rm instr_rm = ( if (((instr_rm = RM_DYN))) then read_reg fcsr_ref \ ((\ (w__0 :: Fcsr) . encdec_rounding_mode_backwards ((get_Fcsr_FRM w__0 :: 3 Word.word)))) else return instr_rm )\ for instr_rm :: " rounding_mode " \ \\val nxFlag : unit -> mword ty5\\ definition nxFlag :: \ unit \(5)Word.word \ where \ nxFlag _ = ( ( 0b00001 :: 5 Word.word))\ \ \\val ufFlag : unit -> mword ty5\\ definition ufFlag :: \ unit \(5)Word.word \ where \ ufFlag _ = ( ( 0b00010 :: 5 Word.word))\ \ \\val ofFlag : unit -> mword ty5\\ definition ofFlag :: \ unit \(5)Word.word \ where \ ofFlag _ = ( ( 0b00100 :: 5 Word.word))\ \ \\val dzFlag : unit -> mword ty5\\ definition dzFlag :: \ unit \(5)Word.word \ where \ dzFlag _ = ( ( 0b01000 :: 5 Word.word))\ \ \\val nvFlag : unit -> mword ty5\\ definition nvFlag :: \ unit \(5)Word.word \ where \ nvFlag _ = ( ( 0b10000 :: 5 Word.word))\ \ \\val fsplit_S : mword ty32 -> (mword ty1 * mword ty8 * mword ty23)\\ definition fsplit_S :: \(32)Word.word \(1)Word.word*(8)Word.word*(23)Word.word \ where \ fsplit_S x32 = ( ((subrange_vec_dec x32 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word), (subrange_vec_dec x32 (( 30 :: int)::ii) (( 23 :: int)::ii) :: 8 Word.word), (subrange_vec_dec x32 (( 22 :: int)::ii) (( 0 :: int)::ii) :: 23 Word.word)))\ for x32 :: "(32)Word.word " \ \\val fmake_S : mword ty1 -> mword ty8 -> mword ty23 -> mword ty32\\ definition fmake_S :: \(1)Word.word \(8)Word.word \(23)Word.word \(32)Word.word \ where \ fmake_S sign exp1 mant = ( (concat_vec sign ((concat_vec exp1 mant :: 31 Word.word)) :: 32 Word.word))\ for sign :: "(1)Word.word " and exp1 :: "(8)Word.word " and mant :: "(23)Word.word " \ \\val canonical_NaN_S : unit -> mword ty32\\ definition canonical_NaN_S :: \ unit \(32)Word.word \ where \ canonical_NaN_S _ = ( ( 0x7FC00000 :: 32 Word.word))\ \ \\val f_is_neg_inf_S : mword ty32 -> bool\\ definition f_is_neg_inf_S :: \(32)Word.word \ bool \ where \ f_is_neg_inf_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((sign = ( 0b1 :: 1 Word.word)))) \ ((((((exp1 = ((ones (( 8 :: int)::ii) :: 8 Word.word))))) \ (((mant = ((zeros_implicit (( 23 :: int)::ii) :: 23 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_neg_norm_S : mword ty32 -> bool\\ definition f_is_neg_norm_S :: \(32)Word.word \ bool \ where \ f_is_neg_norm_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((sign = ( 0b1 :: 1 Word.word)))) \ ((((((exp1 \ ((zeros_implicit (( 8 :: int)::ii) :: 8 Word.word))))) \ (((exp1 \ ((ones (( 8 :: int)::ii) :: 8 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_neg_subnorm_S : mword ty32 -> bool\\ definition f_is_neg_subnorm_S :: \(32)Word.word \ bool \ where \ f_is_neg_subnorm_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((sign = ( 0b1 :: 1 Word.word)))) \ ((((((exp1 = ((zeros_implicit (( 8 :: int)::ii) :: 8 Word.word))))) \ (((mant \ ((zeros_implicit (( 23 :: int)::ii) :: 23 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_neg_zero_S : mword ty32 -> bool\\ definition f_is_neg_zero_S :: \(32)Word.word \ bool \ where \ f_is_neg_zero_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((sign = ((ones (( 1 :: int)::ii) :: 1 Word.word))))) \ ((((((exp1 = ((zeros_implicit (( 8 :: int)::ii) :: 8 Word.word))))) \ (((mant = ((zeros_implicit (( 23 :: int)::ii) :: 23 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_pos_zero_S : mword ty32 -> bool\\ definition f_is_pos_zero_S :: \(32)Word.word \ bool \ where \ f_is_pos_zero_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((sign = ((zeros_implicit (( 1 :: int)::ii) :: 1 Word.word))))) \ ((((((exp1 = ((zeros_implicit (( 8 :: int)::ii) :: 8 Word.word))))) \ (((mant = ((zeros_implicit (( 23 :: int)::ii) :: 23 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_pos_subnorm_S : mword ty32 -> bool\\ definition f_is_pos_subnorm_S :: \(32)Word.word \ bool \ where \ f_is_pos_subnorm_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((sign = ((zeros_implicit (( 1 :: int)::ii) :: 1 Word.word))))) \ ((((((exp1 = ((zeros_implicit (( 8 :: int)::ii) :: 8 Word.word))))) \ (((mant \ ((zeros_implicit (( 23 :: int)::ii) :: 23 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_pos_norm_S : mword ty32 -> bool\\ definition f_is_pos_norm_S :: \(32)Word.word \ bool \ where \ f_is_pos_norm_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((sign = ((zeros_implicit (( 1 :: int)::ii) :: 1 Word.word))))) \ ((((((exp1 \ ((zeros_implicit (( 8 :: int)::ii) :: 8 Word.word))))) \ (((exp1 \ ((ones (( 8 :: int)::ii) :: 8 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_pos_inf_S : mword ty32 -> bool\\ definition f_is_pos_inf_S :: \(32)Word.word \ bool \ where \ f_is_pos_inf_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((sign = ((zeros_implicit (( 1 :: int)::ii) :: 1 Word.word))))) \ ((((((exp1 = ((ones (( 8 :: int)::ii) :: 8 Word.word))))) \ (((mant = ((zeros_implicit (( 23 :: int)::ii) :: 23 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_SNaN_S : mword ty32 -> bool\\ definition f_is_SNaN_S :: \(32)Word.word \ bool \ where \ f_is_SNaN_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((exp1 = ((ones (( 8 :: int)::ii) :: 8 Word.word))))) \ ((((((((access_vec_dec mant (( 22 :: int)::ii))) = B0))) \ (((mant \ ((zeros_implicit (( 23 :: int)::ii) :: 23 Word.word)))))))))))\ for x32 :: "(32)Word.word " \ \\val f_is_QNaN_S : mword ty32 -> bool\\ definition f_is_QNaN_S :: \(32)Word.word \ bool \ where \ f_is_QNaN_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((exp1 = ((ones (( 8 :: int)::ii) :: 8 Word.word))))) \ (((((access_vec_dec mant (( 22 :: int)::ii))) = B1))))))\ for x32 :: "(32)Word.word " \ \\val f_is_NaN_S : mword ty32 -> bool\\ definition f_is_NaN_S :: \(32)Word.word \ bool \ where \ f_is_NaN_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in ((((exp1 = ((ones (( 8 :: int)::ii) :: 8 Word.word))))) \ (((mant \ ((zeros_implicit (( 23 :: int)::ii) :: 23 Word.word))))))))\ for x32 :: "(32)Word.word " \ \\val negate_S : mword ty32 -> mword ty32\\ definition negate_S :: \(32)Word.word \(32)Word.word \ where \ negate_S x32 = ( (let (sign, exp1, mant) = ((fsplit_S x32 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let new_sign = (if (((sign = ( 0b0 :: 1 Word.word)))) then ( 0b1 :: 1 Word.word) else ( 0b0 :: 1 Word.word)) in (fmake_S new_sign exp1 mant :: 32 Word.word))))\ for x32 :: "(32)Word.word " \ \\val feq_quiet_S : mword ty32 -> mword ty32 -> (bool * mword ty5)\\ definition feq_quiet_S :: \(32)Word.word \(32)Word.word \ bool*(5)Word.word \ where \ feq_quiet_S v1 v2 = ( (let (s1, e1, m1) = ((fsplit_S v1 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let (s2, e2, m2) = ((fsplit_S v2 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let v1Is0 = (((f_is_neg_zero_S v1)) \ ((f_is_pos_zero_S v1))) in (let v2Is0 = (((f_is_neg_zero_S v2)) \ ((f_is_pos_zero_S v2))) in (let result = ((((v1 = v2))) \ (((v1Is0 \ v2Is0)))) in (let fflags = (if (((((f_is_SNaN_S v1)) \ ((f_is_SNaN_S v2))))) then (nvFlag () :: 5 Word.word) else (zeros_implicit (( 5 :: int)::ii) :: 5 Word.word)) in (result, fflags))))))))\ for v1 :: "(32)Word.word " and v2 :: "(32)Word.word " \ \\val flt_S : mword ty32 -> mword ty32 -> bool -> (bool * mword ty5)\\ definition flt_S :: \(32)Word.word \(32)Word.word \ bool \ bool*(5)Word.word \ where \ flt_S v1 v2 is_quiet = ( (let (s1, e1, m1) = ((fsplit_S v1 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let (s2, e2, m2) = ((fsplit_S v2 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let (result :: bool) = (if ((((((s1 = ( 0b0 :: 1 Word.word)))) \ (((s2 = ( 0b0 :: 1 Word.word))))))) then if (((e1 = e2))) then ((Word.uint m1)) < ((Word.uint m2)) else ((Word.uint e1)) < ((Word.uint e2)) else if ((((((s1 = ( 0b0 :: 1 Word.word)))) \ (((s2 = ( 0b1 :: 1 Word.word))))))) then False else if ((((((s1 = ( 0b1 :: 1 Word.word)))) \ (((s2 = ( 0b0 :: 1 Word.word))))))) then True else if (((e1 = e2))) then ((Word.uint m1)) > ((Word.uint m2)) else ((Word.uint e1)) > ((Word.uint e2))) in (let fflags = (if is_quiet then if (((((f_is_SNaN_S v1)) \ ((f_is_SNaN_S v2))))) then (nvFlag () :: 5 Word.word) else (zeros_implicit (( 5 :: int)::ii) :: 5 Word.word) else if (((((f_is_NaN_S v1)) \ ((f_is_NaN_S v2))))) then (nvFlag () :: 5 Word.word) else (zeros_implicit (( 5 :: int)::ii) :: 5 Word.word)) in (result, fflags))))))\ for v1 :: "(32)Word.word " and v2 :: "(32)Word.word " and is_quiet :: " bool " \ \\val fle_S : mword ty32 -> mword ty32 -> bool -> (bool * mword ty5)\\ definition fle_S :: \(32)Word.word \(32)Word.word \ bool \ bool*(5)Word.word \ where \ fle_S v1 v2 is_quiet = ( (let (s1, e1, m1) = ((fsplit_S v1 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let (s2, e2, m2) = ((fsplit_S v2 :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let v1Is0 = (((f_is_neg_zero_S v1)) \ ((f_is_pos_zero_S v1))) in (let v2Is0 = (((f_is_neg_zero_S v2)) \ ((f_is_pos_zero_S v2))) in (let (result :: bool) = (if ((((((s1 = ( 0b0 :: 1 Word.word)))) \ (((s2 = ( 0b0 :: 1 Word.word))))))) then if (((e1 = e2))) then ((Word.uint m1)) \ ((Word.uint m2)) else ((Word.uint e1)) < ((Word.uint e2)) else if ((((((s1 = ( 0b0 :: 1 Word.word)))) \ (((s2 = ( 0b1 :: 1 Word.word))))))) then (v1Is0 \ v2Is0) else if ((((((s1 = ( 0b1 :: 1 Word.word)))) \ (((s2 = ( 0b0 :: 1 Word.word))))))) then True else if (((e1 = e2))) then ((Word.uint m1)) \ ((Word.uint m2)) else ((Word.uint e1)) > ((Word.uint e2))) in (let fflags = (if is_quiet then if (((((f_is_SNaN_S v1)) \ ((f_is_SNaN_S v2))))) then (nvFlag () :: 5 Word.word) else (zeros_implicit (( 5 :: int)::ii) :: 5 Word.word) else if (((((f_is_NaN_S v1)) \ ((f_is_NaN_S v2))))) then (nvFlag () :: 5 Word.word) else (zeros_implicit (( 5 :: int)::ii) :: 5 Word.word)) in (result, fflags))))))))\ for v1 :: "(32)Word.word " and v2 :: "(32)Word.word " and is_quiet :: " bool " \ \\val nan_box : mword ty32 -> mword ty32\\ definition nan_box :: \(32)Word.word \(32)Word.word \ where \ nan_box val_32b = ( val_32b )\ for val_32b :: "(32)Word.word " \ \\val nan_unbox : mword ty32 -> mword ty32\\ definition nan_unbox :: \(32)Word.word \(32)Word.word \ where \ nan_unbox regval = ( regval )\ for regval :: "(32)Word.word " \ \\val is_RV32F_or_RV64F : unit -> M bool\\ definition is_RV32F_or_RV64F :: \ unit \((register_value),(bool),(exception))monad \ where \ is_RV32F_or_RV64F _ = ( and_boolM ((haveFExt () )) (return (((((((( 32 :: int)::ii) = (( 32 :: int)::ii)))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))))))))\ \ \\val is_RV64F : unit -> M bool\\ definition is_RV64F :: \ unit \((register_value),(bool),(exception))monad \ where \ is_RV64F _ = ( and_boolM ((haveFExt () )) (return ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))\ \ \\val is_RV32D_or_RV64D : unit -> M bool\\ definition is_RV32D_or_RV64D :: \ unit \((register_value),(bool),(exception))monad \ where \ is_RV32D_or_RV64D _ = ( and_boolM ((haveDExt () )) (return (((((((( 32 :: int)::ii) = (( 32 :: int)::ii)))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))))))))\ \ \\val is_RV64D : unit -> M bool\\ definition is_RV64D :: \ unit \((register_value),(bool),(exception))monad \ where \ is_RV64D _ = ( and_boolM ((haveDExt () )) (return ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))\ \ \\val process_fload64 : mword ty5 -> mword ty32 -> MemoryOpResult (mword ty64) -> M Retired\\ definition process_fload64 :: \(5)Word.word \(32)Word.word \((64)Word.word)MemoryOpResult \((register_value),(Retired),(exception))monad \ where \ process_fload64 rd addr value1 = ( return RETIRE_FAIL )\ for rd :: "(5)Word.word " and addr :: "(32)Word.word " and value1 :: "((64)Word.word)MemoryOpResult " \ \\val process_fload32 : mword ty5 -> mword ty32 -> MemoryOpResult (mword ty32) -> M Retired\\ fun process_fload32 :: \(5)Word.word \(32)Word.word \((32)Word.word)MemoryOpResult \((register_value),(Retired),(exception))monad \ where \ process_fload32 rd addr (MemValue (result)) = ( wF_bits rd ((nan_box result :: 32 Word.word)) \ return RETIRE_SUCCESS )\ for rd :: "(5)Word.word " and addr :: "(32)Word.word " and result :: "(32)Word.word " |\ process_fload32 rd addr (MemException (e)) = ( handle_mem_exception addr e \ return RETIRE_FAIL )\ for rd :: "(5)Word.word " and addr :: "(32)Word.word " and e :: " ExceptionType " \ \\val process_fstore : mword ty32 -> MemoryOpResult bool -> M Retired\\ fun process_fstore :: \(32)Word.word \(bool)MemoryOpResult \((register_value),(Retired),(exception))monad \ where \ process_fstore vaddr (MemValue (True)) = ( return RETIRE_SUCCESS )\ for vaddr :: "(32)Word.word " |\ process_fstore vaddr (MemValue (False)) = ( internal_error (''store got false from mem_write_value''))\ for vaddr :: "(32)Word.word " |\ process_fstore vaddr (MemException (e)) = ( handle_mem_exception vaddr e \ return RETIRE_FAIL )\ for vaddr :: "(32)Word.word " and e :: " ExceptionType " \ \\val f_madd_type_mnemonic_S_forwards : f_madd_op_S -> string\\ \ \\val f_madd_type_mnemonic_S_backwards : string -> M f_madd_op_S\\ \ \\val f_madd_type_mnemonic_S_forwards_matches : f_madd_op_S -> bool\\ \ \\val f_madd_type_mnemonic_S_backwards_matches : string -> bool\\ \ \\val f_madd_type_mnemonic_S_matches_prefix : string -> maybe ((f_madd_op_S * ii))\\ fun f_madd_type_mnemonic_S_forwards :: \ f_madd_op_S \ string \ where \ f_madd_type_mnemonic_S_forwards FMADD_S = ( (''fmadd.s''))\ |\ f_madd_type_mnemonic_S_forwards FMSUB_S = ( (''fmsub.s''))\ |\ f_madd_type_mnemonic_S_forwards FNMSUB_S = ( (''fnmsub.s''))\ |\ f_madd_type_mnemonic_S_forwards FNMADD_S = ( (''fnmadd.s''))\ definition f_madd_type_mnemonic_S_backwards :: \ string \((register_value),(f_madd_op_S),(exception))monad \ where \ f_madd_type_mnemonic_S_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''fmadd.s'')))) then return FMADD_S else if (((p00 = (''fmsub.s'')))) then return FMSUB_S else if (((p00 = (''fnmsub.s'')))) then return FNMSUB_S else if (((p00 = (''fnmadd.s'')))) then return FNMADD_S else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun f_madd_type_mnemonic_S_forwards_matches :: \ f_madd_op_S \ bool \ where \ f_madd_type_mnemonic_S_forwards_matches FMADD_S = ( True )\ |\ f_madd_type_mnemonic_S_forwards_matches FMSUB_S = ( True )\ |\ f_madd_type_mnemonic_S_forwards_matches FNMSUB_S = ( True )\ |\ f_madd_type_mnemonic_S_forwards_matches FNMADD_S = ( True )\ definition f_madd_type_mnemonic_S_backwards_matches :: \ string \ bool \ where \ f_madd_type_mnemonic_S_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''fmadd.s'')))) then True else if (((p00 = (''fmsub.s'')))) then True else if (((p00 = (''fnmsub.s'')))) then True else if (((p00 = (''fnmadd.s'')))) then True else False))\ for arg1 :: " string " \ \\val _s1114_ : string -> maybe string\\ definition s1114 :: \ string \(string)option \ where \ s1114 s11150 = ( (let s11160 = s11150 in if ((string_startswith s11160 (''fnmadd.s''))) then (case ((string_drop s11160 ((string_length (''fnmadd.s''))))) of s1 => Some s1 ) else None))\ for s11150 :: " string " \ \\val _s1110_ : string -> maybe string\\ definition s1110 :: \ string \(string)option \ where \ s1110 s11110 = ( (let s11120 = s11110 in if ((string_startswith s11120 (''fnmsub.s''))) then (case ((string_drop s11120 ((string_length (''fnmsub.s''))))) of s1 => Some s1 ) else None))\ for s11110 :: " string " \ \\val _s1106_ : string -> maybe string\\ definition s1106 :: \ string \(string)option \ where \ s1106 s11070 = ( (let s11080 = s11070 in if ((string_startswith s11080 (''fmsub.s''))) then (case ((string_drop s11080 ((string_length (''fmsub.s''))))) of s1 => Some s1 ) else None))\ for s11070 :: " string " \ \\val _s1102_ : string -> maybe string\\ definition s1102 :: \ string \(string)option \ where \ s1102 s11030 = ( (let s11040 = s11030 in if ((string_startswith s11040 (''fmadd.s''))) then (case ((string_drop s11040 ((string_length (''fmadd.s''))))) of s1 => Some s1 ) else None))\ for s11030 :: " string " definition f_madd_type_mnemonic_S_matches_prefix :: \ string \(f_madd_op_S*int)option \ where \ f_madd_type_mnemonic_S_matches_prefix arg1 = ( (let s11050 = arg1 in if ((case ((s1102 s11050)) of Some (s1) => True | _ => False )) then (case s1102 s11050 of (Some (s1)) => Some (FMADD_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1106 s11050)) of Some (s1) => True | _ => False )) then (case s1106 s11050 of (Some (s1)) => Some (FMSUB_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1110 s11050)) of Some (s1) => True | _ => False )) then (case s1110 s11050 of (Some (s1)) => Some (FNMSUB_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1114 s11050)) of Some (s1) => True | _ => False )) then (case s1114 s11050 of (Some (s1)) => Some (FNMADD_S, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val f_bin_rm_type_mnemonic_S_forwards : f_bin_rm_op_S -> string\\ \ \\val f_bin_rm_type_mnemonic_S_backwards : string -> M f_bin_rm_op_S\\ \ \\val f_bin_rm_type_mnemonic_S_forwards_matches : f_bin_rm_op_S -> bool\\ \ \\val f_bin_rm_type_mnemonic_S_backwards_matches : string -> bool\\ \ \\val f_bin_rm_type_mnemonic_S_matches_prefix : string -> maybe ((f_bin_rm_op_S * ii))\\ fun f_bin_rm_type_mnemonic_S_forwards :: \ f_bin_rm_op_S \ string \ where \ f_bin_rm_type_mnemonic_S_forwards FADD_S = ( (''fadd.s''))\ |\ f_bin_rm_type_mnemonic_S_forwards FSUB_S = ( (''fsub.s''))\ |\ f_bin_rm_type_mnemonic_S_forwards FMUL_S = ( (''fmul.s''))\ |\ f_bin_rm_type_mnemonic_S_forwards FDIV_S = ( (''fdiv.s''))\ definition f_bin_rm_type_mnemonic_S_backwards :: \ string \((register_value),(f_bin_rm_op_S),(exception))monad \ where \ f_bin_rm_type_mnemonic_S_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''fadd.s'')))) then return FADD_S else if (((p00 = (''fsub.s'')))) then return FSUB_S else if (((p00 = (''fmul.s'')))) then return FMUL_S else if (((p00 = (''fdiv.s'')))) then return FDIV_S else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun f_bin_rm_type_mnemonic_S_forwards_matches :: \ f_bin_rm_op_S \ bool \ where \ f_bin_rm_type_mnemonic_S_forwards_matches FADD_S = ( True )\ |\ f_bin_rm_type_mnemonic_S_forwards_matches FSUB_S = ( True )\ |\ f_bin_rm_type_mnemonic_S_forwards_matches FMUL_S = ( True )\ |\ f_bin_rm_type_mnemonic_S_forwards_matches FDIV_S = ( True )\ definition f_bin_rm_type_mnemonic_S_backwards_matches :: \ string \ bool \ where \ f_bin_rm_type_mnemonic_S_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''fadd.s'')))) then True else if (((p00 = (''fsub.s'')))) then True else if (((p00 = (''fmul.s'')))) then True else if (((p00 = (''fdiv.s'')))) then True else False))\ for arg1 :: " string " \ \\val _s1130_ : string -> maybe string\\ definition s1130 :: \ string \(string)option \ where \ s1130 s11310 = ( (let s11320 = s11310 in if ((string_startswith s11320 (''fdiv.s''))) then (case ((string_drop s11320 ((string_length (''fdiv.s''))))) of s1 => Some s1 ) else None))\ for s11310 :: " string " \ \\val _s1126_ : string -> maybe string\\ definition s1126 :: \ string \(string)option \ where \ s1126 s11270 = ( (let s11280 = s11270 in if ((string_startswith s11280 (''fmul.s''))) then (case ((string_drop s11280 ((string_length (''fmul.s''))))) of s1 => Some s1 ) else None))\ for s11270 :: " string " \ \\val _s1122_ : string -> maybe string\\ definition s1122 :: \ string \(string)option \ where \ s1122 s11230 = ( (let s11240 = s11230 in if ((string_startswith s11240 (''fsub.s''))) then (case ((string_drop s11240 ((string_length (''fsub.s''))))) of s1 => Some s1 ) else None))\ for s11230 :: " string " \ \\val _s1118_ : string -> maybe string\\ definition s1118 :: \ string \(string)option \ where \ s1118 s11190 = ( (let s11200 = s11190 in if ((string_startswith s11200 (''fadd.s''))) then (case ((string_drop s11200 ((string_length (''fadd.s''))))) of s1 => Some s1 ) else None))\ for s11190 :: " string " definition f_bin_rm_type_mnemonic_S_matches_prefix :: \ string \(f_bin_rm_op_S*int)option \ where \ f_bin_rm_type_mnemonic_S_matches_prefix arg1 = ( (let s11210 = arg1 in if ((case ((s1118 s11210)) of Some (s1) => True | _ => False )) then (case s1118 s11210 of (Some (s1)) => Some (FADD_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1122 s11210)) of Some (s1) => True | _ => False )) then (case s1122 s11210 of (Some (s1)) => Some (FSUB_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1126 s11210)) of Some (s1) => True | _ => False )) then (case s1126 s11210 of (Some (s1)) => Some (FMUL_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1130 s11210)) of Some (s1) => True | _ => False )) then (case s1130 s11210 of (Some (s1)) => Some (FDIV_S, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val f_un_rm_type_mnemonic_S_forwards : f_un_rm_op_S -> string\\ \ \\val f_un_rm_type_mnemonic_S_backwards : string -> M f_un_rm_op_S\\ \ \\val f_un_rm_type_mnemonic_S_forwards_matches : f_un_rm_op_S -> bool\\ \ \\val f_un_rm_type_mnemonic_S_backwards_matches : string -> bool\\ \ \\val f_un_rm_type_mnemonic_S_matches_prefix : string -> maybe ((f_un_rm_op_S * ii))\\ fun f_un_rm_type_mnemonic_S_forwards :: \ f_un_rm_op_S \ string \ where \ f_un_rm_type_mnemonic_S_forwards FSQRT_S = ( (''fsqrt.s''))\ |\ f_un_rm_type_mnemonic_S_forwards FCVT_W_S = ( (''fcvt.w.s''))\ |\ f_un_rm_type_mnemonic_S_forwards FCVT_WU_S = ( (''fcvt.wu.s''))\ |\ f_un_rm_type_mnemonic_S_forwards FCVT_S_W = ( (''fcvt.s.w''))\ |\ f_un_rm_type_mnemonic_S_forwards FCVT_S_WU = ( (''fcvt.s.wu''))\ |\ f_un_rm_type_mnemonic_S_forwards FCVT_L_S = ( (''fcvt.l.s''))\ |\ f_un_rm_type_mnemonic_S_forwards FCVT_LU_S = ( (''fcvt.lu.s''))\ |\ f_un_rm_type_mnemonic_S_forwards FCVT_S_L = ( (''fcvt.s.l''))\ |\ f_un_rm_type_mnemonic_S_forwards FCVT_S_LU = ( (''fcvt.s.lu''))\ definition f_un_rm_type_mnemonic_S_backwards :: \ string \((register_value),(f_un_rm_op_S),(exception))monad \ where \ f_un_rm_type_mnemonic_S_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''fsqrt.s'')))) then return FSQRT_S else if (((p00 = (''fcvt.w.s'')))) then return FCVT_W_S else if (((p00 = (''fcvt.wu.s'')))) then return FCVT_WU_S else if (((p00 = (''fcvt.s.w'')))) then return FCVT_S_W else if (((p00 = (''fcvt.s.wu'')))) then return FCVT_S_WU else if (((p00 = (''fcvt.l.s'')))) then return FCVT_L_S else if (((p00 = (''fcvt.lu.s'')))) then return FCVT_LU_S else if (((p00 = (''fcvt.s.l'')))) then return FCVT_S_L else if (((p00 = (''fcvt.s.lu'')))) then return FCVT_S_LU else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun f_un_rm_type_mnemonic_S_forwards_matches :: \ f_un_rm_op_S \ bool \ where \ f_un_rm_type_mnemonic_S_forwards_matches FSQRT_S = ( True )\ |\ f_un_rm_type_mnemonic_S_forwards_matches FCVT_W_S = ( True )\ |\ f_un_rm_type_mnemonic_S_forwards_matches FCVT_WU_S = ( True )\ |\ f_un_rm_type_mnemonic_S_forwards_matches FCVT_S_W = ( True )\ |\ f_un_rm_type_mnemonic_S_forwards_matches FCVT_S_WU = ( True )\ |\ f_un_rm_type_mnemonic_S_forwards_matches FCVT_L_S = ( True )\ |\ f_un_rm_type_mnemonic_S_forwards_matches FCVT_LU_S = ( True )\ |\ f_un_rm_type_mnemonic_S_forwards_matches FCVT_S_L = ( True )\ |\ f_un_rm_type_mnemonic_S_forwards_matches FCVT_S_LU = ( True )\ definition f_un_rm_type_mnemonic_S_backwards_matches :: \ string \ bool \ where \ f_un_rm_type_mnemonic_S_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''fsqrt.s'')))) then True else if (((p00 = (''fcvt.w.s'')))) then True else if (((p00 = (''fcvt.wu.s'')))) then True else if (((p00 = (''fcvt.s.w'')))) then True else if (((p00 = (''fcvt.s.wu'')))) then True else if (((p00 = (''fcvt.l.s'')))) then True else if (((p00 = (''fcvt.lu.s'')))) then True else if (((p00 = (''fcvt.s.l'')))) then True else if (((p00 = (''fcvt.s.lu'')))) then True else False))\ for arg1 :: " string " \ \\val _s1166_ : string -> maybe string\\ definition s1166 :: \ string \(string)option \ where \ s1166 s11670 = ( (let s11680 = s11670 in if ((string_startswith s11680 (''fcvt.s.lu''))) then (case ((string_drop s11680 ((string_length (''fcvt.s.lu''))))) of s1 => Some s1 ) else None))\ for s11670 :: " string " \ \\val _s1162_ : string -> maybe string\\ definition s1162 :: \ string \(string)option \ where \ s1162 s11630 = ( (let s11640 = s11630 in if ((string_startswith s11640 (''fcvt.s.l''))) then (case ((string_drop s11640 ((string_length (''fcvt.s.l''))))) of s1 => Some s1 ) else None))\ for s11630 :: " string " \ \\val _s1158_ : string -> maybe string\\ definition s1158 :: \ string \(string)option \ where \ s1158 s11590 = ( (let s11600 = s11590 in if ((string_startswith s11600 (''fcvt.lu.s''))) then (case ((string_drop s11600 ((string_length (''fcvt.lu.s''))))) of s1 => Some s1 ) else None))\ for s11590 :: " string " \ \\val _s1154_ : string -> maybe string\\ definition s1154 :: \ string \(string)option \ where \ s1154 s11550 = ( (let s11560 = s11550 in if ((string_startswith s11560 (''fcvt.l.s''))) then (case ((string_drop s11560 ((string_length (''fcvt.l.s''))))) of s1 => Some s1 ) else None))\ for s11550 :: " string " \ \\val _s1150_ : string -> maybe string\\ definition s1150 :: \ string \(string)option \ where \ s1150 s11510 = ( (let s11520 = s11510 in if ((string_startswith s11520 (''fcvt.s.wu''))) then (case ((string_drop s11520 ((string_length (''fcvt.s.wu''))))) of s1 => Some s1 ) else None))\ for s11510 :: " string " \ \\val _s1146_ : string -> maybe string\\ definition s1146 :: \ string \(string)option \ where \ s1146 s11470 = ( (let s11480 = s11470 in if ((string_startswith s11480 (''fcvt.s.w''))) then (case ((string_drop s11480 ((string_length (''fcvt.s.w''))))) of s1 => Some s1 ) else None))\ for s11470 :: " string " \ \\val _s1142_ : string -> maybe string\\ definition s1142 :: \ string \(string)option \ where \ s1142 s11430 = ( (let s11440 = s11430 in if ((string_startswith s11440 (''fcvt.wu.s''))) then (case ((string_drop s11440 ((string_length (''fcvt.wu.s''))))) of s1 => Some s1 ) else None))\ for s11430 :: " string " \ \\val _s1138_ : string -> maybe string\\ definition s1138 :: \ string \(string)option \ where \ s1138 s11390 = ( (let s11400 = s11390 in if ((string_startswith s11400 (''fcvt.w.s''))) then (case ((string_drop s11400 ((string_length (''fcvt.w.s''))))) of s1 => Some s1 ) else None))\ for s11390 :: " string " \ \\val _s1134_ : string -> maybe string\\ definition s1134 :: \ string \(string)option \ where \ s1134 s11350 = ( (let s11360 = s11350 in if ((string_startswith s11360 (''fsqrt.s''))) then (case ((string_drop s11360 ((string_length (''fsqrt.s''))))) of s1 => Some s1 ) else None))\ for s11350 :: " string " definition f_un_rm_type_mnemonic_S_matches_prefix :: \ string \(f_un_rm_op_S*int)option \ where \ f_un_rm_type_mnemonic_S_matches_prefix arg1 = ( (let s11370 = arg1 in if ((case ((s1134 s11370)) of Some (s1) => True | _ => False )) then (case s1134 s11370 of (Some (s1)) => Some (FSQRT_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1138 s11370)) of Some (s1) => True | _ => False )) then (case s1138 s11370 of (Some (s1)) => Some (FCVT_W_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1142 s11370)) of Some (s1) => True | _ => False )) then (case s1142 s11370 of (Some (s1)) => Some (FCVT_WU_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1146 s11370)) of Some (s1) => True | _ => False )) then (case s1146 s11370 of (Some (s1)) => Some (FCVT_S_W, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1150 s11370)) of Some (s1) => True | _ => False )) then (case s1150 s11370 of (Some (s1)) => Some (FCVT_S_WU, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1154 s11370)) of Some (s1) => True | _ => False )) then (case s1154 s11370 of (Some (s1)) => Some (FCVT_L_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1158 s11370)) of Some (s1) => True | _ => False )) then (case s1158 s11370 of (Some (s1)) => Some (FCVT_LU_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1162 s11370)) of Some (s1) => True | _ => False )) then (case s1162 s11370 of (Some (s1)) => Some (FCVT_S_L, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1166 s11370)) of Some (s1) => True | _ => False )) then (case s1166 s11370 of (Some (s1)) => Some (FCVT_S_LU, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val f_bin_type_mnemonic_S_forwards : f_bin_op_S -> string\\ \ \\val f_bin_type_mnemonic_S_backwards : string -> M f_bin_op_S\\ \ \\val f_bin_type_mnemonic_S_forwards_matches : f_bin_op_S -> bool\\ \ \\val f_bin_type_mnemonic_S_backwards_matches : string -> bool\\ \ \\val f_bin_type_mnemonic_S_matches_prefix : string -> maybe ((f_bin_op_S * ii))\\ fun f_bin_type_mnemonic_S_forwards :: \ f_bin_op_S \ string \ where \ f_bin_type_mnemonic_S_forwards FSGNJ_S = ( (''fsgnj.s''))\ |\ f_bin_type_mnemonic_S_forwards FSGNJN_S = ( (''fsgnjn.s''))\ |\ f_bin_type_mnemonic_S_forwards FSGNJX_S = ( (''fsgnjx.s''))\ |\ f_bin_type_mnemonic_S_forwards FMIN_S = ( (''fmin.s''))\ |\ f_bin_type_mnemonic_S_forwards FMAX_S = ( (''fmax.s''))\ |\ f_bin_type_mnemonic_S_forwards FEQ_S = ( (''feq.s''))\ |\ f_bin_type_mnemonic_S_forwards FLT_S = ( (''flt.s''))\ |\ f_bin_type_mnemonic_S_forwards FLE_S = ( (''fle.s''))\ definition f_bin_type_mnemonic_S_backwards :: \ string \((register_value),(f_bin_op_S),(exception))monad \ where \ f_bin_type_mnemonic_S_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''fsgnj.s'')))) then return FSGNJ_S else if (((p00 = (''fsgnjn.s'')))) then return FSGNJN_S else if (((p00 = (''fsgnjx.s'')))) then return FSGNJX_S else if (((p00 = (''fmin.s'')))) then return FMIN_S else if (((p00 = (''fmax.s'')))) then return FMAX_S else if (((p00 = (''feq.s'')))) then return FEQ_S else if (((p00 = (''flt.s'')))) then return FLT_S else if (((p00 = (''fle.s'')))) then return FLE_S else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun f_bin_type_mnemonic_S_forwards_matches :: \ f_bin_op_S \ bool \ where \ f_bin_type_mnemonic_S_forwards_matches FSGNJ_S = ( True )\ |\ f_bin_type_mnemonic_S_forwards_matches FSGNJN_S = ( True )\ |\ f_bin_type_mnemonic_S_forwards_matches FSGNJX_S = ( True )\ |\ f_bin_type_mnemonic_S_forwards_matches FMIN_S = ( True )\ |\ f_bin_type_mnemonic_S_forwards_matches FMAX_S = ( True )\ |\ f_bin_type_mnemonic_S_forwards_matches FEQ_S = ( True )\ |\ f_bin_type_mnemonic_S_forwards_matches FLT_S = ( True )\ |\ f_bin_type_mnemonic_S_forwards_matches FLE_S = ( True )\ definition f_bin_type_mnemonic_S_backwards_matches :: \ string \ bool \ where \ f_bin_type_mnemonic_S_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''fsgnj.s'')))) then True else if (((p00 = (''fsgnjn.s'')))) then True else if (((p00 = (''fsgnjx.s'')))) then True else if (((p00 = (''fmin.s'')))) then True else if (((p00 = (''fmax.s'')))) then True else if (((p00 = (''feq.s'')))) then True else if (((p00 = (''flt.s'')))) then True else if (((p00 = (''fle.s'')))) then True else False))\ for arg1 :: " string " \ \\val _s1198_ : string -> maybe string\\ definition s1198 :: \ string \(string)option \ where \ s1198 s11990 = ( (let s12000 = s11990 in if ((string_startswith s12000 (''fle.s''))) then (case ((string_drop s12000 ((string_length (''fle.s''))))) of s1 => Some s1 ) else None))\ for s11990 :: " string " \ \\val _s1194_ : string -> maybe string\\ definition s1194 :: \ string \(string)option \ where \ s1194 s11950 = ( (let s11960 = s11950 in if ((string_startswith s11960 (''flt.s''))) then (case ((string_drop s11960 ((string_length (''flt.s''))))) of s1 => Some s1 ) else None))\ for s11950 :: " string " \ \\val _s1190_ : string -> maybe string\\ definition s1190 :: \ string \(string)option \ where \ s1190 s11910 = ( (let s11920 = s11910 in if ((string_startswith s11920 (''feq.s''))) then (case ((string_drop s11920 ((string_length (''feq.s''))))) of s1 => Some s1 ) else None))\ for s11910 :: " string " \ \\val _s1186_ : string -> maybe string\\ definition s1186 :: \ string \(string)option \ where \ s1186 s11870 = ( (let s11880 = s11870 in if ((string_startswith s11880 (''fmax.s''))) then (case ((string_drop s11880 ((string_length (''fmax.s''))))) of s1 => Some s1 ) else None))\ for s11870 :: " string " \ \\val _s1182_ : string -> maybe string\\ definition s1182 :: \ string \(string)option \ where \ s1182 s11830 = ( (let s11840 = s11830 in if ((string_startswith s11840 (''fmin.s''))) then (case ((string_drop s11840 ((string_length (''fmin.s''))))) of s1 => Some s1 ) else None))\ for s11830 :: " string " \ \\val _s1178_ : string -> maybe string\\ definition s1178 :: \ string \(string)option \ where \ s1178 s11790 = ( (let s11800 = s11790 in if ((string_startswith s11800 (''fsgnjx.s''))) then (case ((string_drop s11800 ((string_length (''fsgnjx.s''))))) of s1 => Some s1 ) else None))\ for s11790 :: " string " \ \\val _s1174_ : string -> maybe string\\ definition s1174 :: \ string \(string)option \ where \ s1174 s11750 = ( (let s11760 = s11750 in if ((string_startswith s11760 (''fsgnjn.s''))) then (case ((string_drop s11760 ((string_length (''fsgnjn.s''))))) of s1 => Some s1 ) else None))\ for s11750 :: " string " \ \\val _s1170_ : string -> maybe string\\ definition s1170 :: \ string \(string)option \ where \ s1170 s11710 = ( (let s11720 = s11710 in if ((string_startswith s11720 (''fsgnj.s''))) then (case ((string_drop s11720 ((string_length (''fsgnj.s''))))) of s1 => Some s1 ) else None))\ for s11710 :: " string " definition f_bin_type_mnemonic_S_matches_prefix :: \ string \(f_bin_op_S*int)option \ where \ f_bin_type_mnemonic_S_matches_prefix arg1 = ( (let s11730 = arg1 in if ((case ((s1170 s11730)) of Some (s1) => True | _ => False )) then (case s1170 s11730 of (Some (s1)) => Some (FSGNJ_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1174 s11730)) of Some (s1) => True | _ => False )) then (case s1174 s11730 of (Some (s1)) => Some (FSGNJN_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1178 s11730)) of Some (s1) => True | _ => False )) then (case s1178 s11730 of (Some (s1)) => Some (FSGNJX_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1182 s11730)) of Some (s1) => True | _ => False )) then (case s1182 s11730 of (Some (s1)) => Some (FMIN_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1186 s11730)) of Some (s1) => True | _ => False )) then (case s1186 s11730 of (Some (s1)) => Some (FMAX_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1190 s11730)) of Some (s1) => True | _ => False )) then (case s1190 s11730 of (Some (s1)) => Some (FEQ_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1194 s11730)) of Some (s1) => True | _ => False )) then (case s1194 s11730 of (Some (s1)) => Some (FLT_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1198 s11730)) of Some (s1) => True | _ => False )) then (case s1198 s11730 of (Some (s1)) => Some (FLE_S, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " \ \\val f_un_type_mnemonic_S_forwards : f_un_op_S -> string\\ \ \\val f_un_type_mnemonic_S_backwards : string -> M f_un_op_S\\ \ \\val f_un_type_mnemonic_S_forwards_matches : f_un_op_S -> bool\\ \ \\val f_un_type_mnemonic_S_backwards_matches : string -> bool\\ \ \\val f_un_type_mnemonic_S_matches_prefix : string -> maybe ((f_un_op_S * ii))\\ fun f_un_type_mnemonic_S_forwards :: \ f_un_op_S \ string \ where \ f_un_type_mnemonic_S_forwards FMV_X_W = ( (''fmv.x.w''))\ |\ f_un_type_mnemonic_S_forwards FCLASS_S = ( (''fclass.s''))\ |\ f_un_type_mnemonic_S_forwards FMV_W_X = ( (''fmv.w.x''))\ definition f_un_type_mnemonic_S_backwards :: \ string \((register_value),(f_un_op_S),(exception))monad \ where \ f_un_type_mnemonic_S_backwards arg1 = ( (let p00 = arg1 in if (((p00 = (''fmv.x.w'')))) then return FMV_X_W else if (((p00 = (''fclass.s'')))) then return FCLASS_S else if (((p00 = (''fmv.w.x'')))) then return FMV_W_X else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " fun f_un_type_mnemonic_S_forwards_matches :: \ f_un_op_S \ bool \ where \ f_un_type_mnemonic_S_forwards_matches FMV_X_W = ( True )\ |\ f_un_type_mnemonic_S_forwards_matches FCLASS_S = ( True )\ |\ f_un_type_mnemonic_S_forwards_matches FMV_W_X = ( True )\ definition f_un_type_mnemonic_S_backwards_matches :: \ string \ bool \ where \ f_un_type_mnemonic_S_backwards_matches arg1 = ( (let p00 = arg1 in if (((p00 = (''fmv.x.w'')))) then True else if (((p00 = (''fclass.s'')))) then True else if (((p00 = (''fmv.w.x'')))) then True else False))\ for arg1 :: " string " \ \\val _s1210_ : string -> maybe string\\ definition s1210 :: \ string \(string)option \ where \ s1210 s12110 = ( (let s12120 = s12110 in if ((string_startswith s12120 (''fmv.w.x''))) then (case ((string_drop s12120 ((string_length (''fmv.w.x''))))) of s1 => Some s1 ) else None))\ for s12110 :: " string " \ \\val _s1206_ : string -> maybe string\\ definition s1206 :: \ string \(string)option \ where \ s1206 s12070 = ( (let s12080 = s12070 in if ((string_startswith s12080 (''fclass.s''))) then (case ((string_drop s12080 ((string_length (''fclass.s''))))) of s1 => Some s1 ) else None))\ for s12070 :: " string " \ \\val _s1202_ : string -> maybe string\\ definition s1202 :: \ string \(string)option \ where \ s1202 s12030 = ( (let s12040 = s12030 in if ((string_startswith s12040 (''fmv.x.w''))) then (case ((string_drop s12040 ((string_length (''fmv.x.w''))))) of s1 => Some s1 ) else None))\ for s12030 :: " string " definition f_un_type_mnemonic_S_matches_prefix :: \ string \(f_un_op_S*int)option \ where \ f_un_type_mnemonic_S_matches_prefix arg1 = ( (let s12050 = arg1 in if ((case ((s1202 s12050)) of Some (s1) => True | _ => False )) then (case s1202 s12050 of (Some (s1)) => Some (FMV_X_W, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1206 s12050)) of Some (s1) => True | _ => False )) then (case s1206 s12050 of (Some (s1)) => Some (FCLASS_S, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s1210 s12050)) of Some (s1) => True | _ => False )) then (case s1210 s12050 of (Some (s1)) => Some (FMV_W_X, ((string_length arg1)) - ((string_length s1))) ) else None))\ for arg1 :: " string " definition encdec_forwards :: \ ast \((register_value),((32)Word.word),(exception))monad \ where \ encdec_forwards ast = ( (let arg1 = ast in (case arg1 of UTYPE ((imm, rd, op1)) => return ((concat_vec imm ((concat_vec rd ((encdec_uop_forwards op1 :: 7 Word.word)) :: 12 Word.word)) :: 32 Word.word)) | RISCV_JAL ((v__2, rd)) => if (((((subrange_vec_dec v__2 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__2 (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word)) in (let (imm_8 :: 1 bits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in (let (imm_7_0 :: 8 bits) = ((subrange_vec_dec v__2 (( 19 :: int)::ii) (( 12 :: int)::ii) :: 8 Word.word)) in (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__2 (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word)) in (let (imm_18_13 :: 6 bits) = ((subrange_vec_dec v__2 (( 10 :: int)::ii) (( 5 :: int)::ii) :: 6 Word.word)) in (let (imm_12_9 :: 4 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 1 :: int)::ii) :: 4 Word.word)) in return ((concat_vec imm_19 ((concat_vec imm_18_13 ((concat_vec imm_12_9 ((concat_vec imm_8 ((concat_vec imm_7_0 ((concat_vec rd ( 0b1101111 :: 7 Word.word) :: 12 Word.word)) :: 20 Word.word)) :: 21 Word.word)) :: 25 Word.word)) :: 31 Word.word)) :: 32 Word.word)))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | RISCV_JALR ((imm, rs1, rd)) => return ((concat_vec imm ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b1100111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) | BTYPE ((v__4, rs2, rs1, op1)) => if (((((subrange_vec_dec v__4 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__4 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__4 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm7_5_0 :: 6 bits) = ((subrange_vec_dec v__4 (( 10 :: int)::ii) (( 5 :: int)::ii) :: 6 Word.word)) in (let (imm5_4_1 :: 4 bits) = ((subrange_vec_dec v__4 (( 4 :: int)::ii) (( 1 :: int)::ii) :: 4 Word.word)) in (let (imm5_0 :: 1 bits) = ((subrange_vec_dec v__4 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in return ((concat_vec imm7_6 ((concat_vec imm7_5_0 ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_bop_forwards op1 :: 3 Word.word)) ((concat_vec imm5_4_1 ((concat_vec imm5_0 ( 0b1100011 :: 7 Word.word) :: 8 Word.word)) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 31 Word.word)) :: 32 Word.word))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | ITYPE ((imm, rs1, rd, op1)) => return ((concat_vec imm ((concat_vec rs1 ((concat_vec ((encdec_iop_forwards op1 :: 3 Word.word)) ((concat_vec rd ( 0b0010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) | SHIFTIOP ((shamt, rs1, rd, RISCV_SLLI)) => return ((concat_vec ( 0b000000 :: 6 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b0010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 26 Word.word)) :: 32 Word.word)) | SHIFTIOP ((shamt, rs1, rd, RISCV_SRLI)) => return ((concat_vec ( 0b000000 :: 6 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 26 Word.word)) :: 32 Word.word)) | SHIFTIOP ((shamt, rs1, rd, RISCV_SRAI)) => return ((concat_vec ( 0b010000 :: 6 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 26 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_ADD)) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_SLT)) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_SLTU)) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_AND)) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b111 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_OR)) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b110 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_XOR)) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_SLL)) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_SRL)) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_SUB)) => return ((concat_vec ( 0b0100000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | RTYPE ((rs2, rs1, rd, RISCV_SRA)) => return ((concat_vec ( 0b0100000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | LOAD ((imm, rs1, rd, is_unsigned, size1, False, False)) => if (((((((word_width_bytes size1)) < (( 4 :: int)::ii))) \ (((((\ is_unsigned)) \ ((((word_width_bytes size1)) \ (( 4 :: int)::ii))))))))) then return ((concat_vec imm ((concat_vec rs1 ((concat_vec ((bool_bits_forwards is_unsigned :: 1 Word.word)) ((concat_vec ((size_bits_forwards size1 :: 2 Word.word)) ((concat_vec rd ( 0b0000011 :: 7 Word.word) :: 12 Word.word)) :: 14 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | STORE ((v__6, rs2, rs1, size1, False, False)) => if ((((word_width_bytes size1)) \ (( 4 :: int)::ii))) then (let (imm7 :: 7 bits) = ((subrange_vec_dec v__6 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in (let (imm7 :: 7 bits) = ((subrange_vec_dec v__6 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in (let (imm5 :: 5 bits) = ((subrange_vec_dec v__6 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec imm7 ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ((size_bits_forwards size1 :: 2 Word.word)) ((concat_vec imm5 ( 0b0100011 :: 7 Word.word) :: 12 Word.word)) :: 14 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | ADDIW ((imm, rs1, rd)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec imm ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b0011011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | SHIFTW ((shamt, rs1, rd, RISCV_SLLI)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b0011011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | SHIFTW ((shamt, rs1, rd, RISCV_SRLI)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0011011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | SHIFTW ((shamt, rs1, rd, RISCV_SRAI)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0100000 :: 7 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0011011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | RTYPEW ((rs2, rs1, rd, RISCV_ADDW)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b0111011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | RTYPEW ((rs2, rs1, rd, RISCV_SUBW)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0100000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b0111011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | RTYPEW ((rs2, rs1, rd, RISCV_SLLW)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b0111011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | RTYPEW ((rs2, rs1, rd, RISCV_SRLW)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0111011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | RTYPEW ((rs2, rs1, rd, RISCV_SRAW)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0100000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0111011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | SHIFTIWOP ((shamt, rs1, rd, RISCV_SLLIW)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b0011011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | SHIFTIWOP ((shamt, rs1, rd, RISCV_SRLIW)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0011011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | SHIFTIWOP ((shamt, rs1, rd, RISCV_SRAIW)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0100000 :: 7 Word.word) ((concat_vec shamt ((concat_vec rs1 ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec rd ( 0b0011011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | FENCE ((pred, succ)) => return ((concat_vec ( 0x0 :: 4 Word.word) ((concat_vec pred ((concat_vec succ ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b0001111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 24 Word.word)) :: 28 Word.word)) :: 32 Word.word)) | FENCE_TSO ((pred, succ)) => return ((concat_vec ( 0x8 :: 4 Word.word) ((concat_vec pred ((concat_vec succ ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b0001111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 24 Word.word)) :: 28 Word.word)) :: 32 Word.word)) | FENCEI (_) => return ((concat_vec ( 0x000 :: 12 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b0001111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) | ECALL (_) => return ((concat_vec ( 0x000 :: 12 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b1110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) | MRET (_) => return ((concat_vec ( 0b0011000 :: 7 Word.word) ((concat_vec ( 0b00010 :: 5 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b1110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | SRET (_) => return ((concat_vec ( 0b0001000 :: 7 Word.word) ((concat_vec ( 0b00010 :: 5 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b1110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | EBREAK (_) => return ((concat_vec ( 0x001 :: 12 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b1110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) | WFI (_) => return ((concat_vec ( 0x105 :: 12 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b1110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) | SFENCE_VMA ((rs1, rs2)) => return ((concat_vec ( 0b0001001 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b1110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | LOADRES ((aq, rl, rs1, size1, rd)) => if ((amo_width_valid size1)) then return ((concat_vec ( 0b00010 :: 5 Word.word) ((concat_vec ((bool_bits_forwards aq :: 1 Word.word)) ((concat_vec ((bool_bits_forwards rl :: 1 Word.word)) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ((size_bits_forwards size1 :: 2 Word.word)) ((concat_vec rd ( 0b0101111 :: 7 Word.word) :: 12 Word.word)) :: 14 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 26 Word.word)) :: 27 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | STORECON ((aq, rl, rs2, rs1, size1, rd)) => if ((amo_width_valid size1)) then return ((concat_vec ( 0b00011 :: 5 Word.word) ((concat_vec ((bool_bits_forwards aq :: 1 Word.word)) ((concat_vec ((bool_bits_forwards rl :: 1 Word.word)) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ((size_bits_forwards size1 :: 2 Word.word)) ((concat_vec rd ( 0b0101111 :: 7 Word.word) :: 12 Word.word)) :: 14 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 26 Word.word)) :: 27 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | AMO ((op1, aq, rl, rs2, rs1, size1, rd)) => if ((amo_width_valid size1)) then return ((concat_vec ((encdec_amoop_forwards op1 :: 5 Word.word)) ((concat_vec ((bool_bits_forwards aq :: 1 Word.word)) ((concat_vec ((bool_bits_forwards rl :: 1 Word.word)) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ((size_bits_forwards size1 :: 2 Word.word)) ((concat_vec rd ( 0b0101111 :: 7 Word.word) :: 12 Word.word)) :: 14 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 26 Word.word)) :: 27 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | MUL ((rs2, rs1, rd, high, signed1, signed2)) => return ((concat_vec ( 0b0000001 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_mul_op_forwards (high, signed1, signed2) :: 3 Word.word)) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | DIV ((rs2, rs1, rd, s)) => return ((concat_vec ( 0b0000001 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b10 :: 2 Word.word) ((concat_vec ((bool_not_bits_forwards s :: 1 Word.word)) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 13 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | REM ((rs2, rs1, rd, s)) => return ((concat_vec ( 0b0000001 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec ((bool_not_bits_forwards s :: 1 Word.word)) ((concat_vec rd ( 0b0110011 :: 7 Word.word) :: 12 Word.word)) :: 13 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | MULW ((rs2, rs1, rd)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000001 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b0111011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | DIVW ((rs2, rs1, rd, s)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000001 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b10 :: 2 Word.word) ((concat_vec ((bool_not_bits_forwards s :: 1 Word.word)) ((concat_vec rd ( 0b0111011 :: 7 Word.word) :: 12 Word.word)) :: 13 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | REMW ((rs2, rs1, rd, s)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b0000001 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec ((bool_not_bits_forwards s :: 1 Word.word)) ((concat_vec rd ( 0b0111011 :: 7 Word.word) :: 12 Word.word)) :: 13 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | CSR ((csr, rs1, rd, is_imm, op1)) => return ((concat_vec csr ((concat_vec rs1 ((concat_vec ((bool_bits_forwards is_imm :: 1 Word.word)) ((concat_vec ((encdec_csrop_forwards op1 :: 2 Word.word)) ((concat_vec rd ( 0b1110011 :: 7 Word.word) :: 12 Word.word)) :: 14 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) | URET (_) => return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec ( 0b00010 :: 5 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b1110011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) | FENCE_RESERVED ((fm, pred, succ, rs, rd)) => if (((((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word)))))))))) then return ((concat_vec fm ((concat_vec pred ((concat_vec succ ((concat_vec rs ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b0001111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 24 Word.word)) :: 28 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | FENCEI_RESERVED ((imm, rs, rd)) => if ((((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg))))))))) then return ((concat_vec imm ((concat_vec rs ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b0001111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | LOAD_FP ((imm, rs1, rd, WORD)) => is_RV32F_or_RV64F () \ ((\ (w__48 :: bool) . if w__48 then return ((concat_vec imm ((concat_vec rs1 ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec rd ( 0b0000111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | LOAD_FP ((imm, rs1, rd, DOUBLE)) => is_RV32D_or_RV64D () \ ((\ (w__51 :: bool) . if w__51 then return ((concat_vec imm ((concat_vec rs1 ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec rd ( 0b0000111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | STORE_FP ((v__7, rs2, rs1, WORD)) => is_RV32F_or_RV64F () \ ((\ (w__54 :: bool) . if w__54 then (let (imm7 :: 7 bits) = ((subrange_vec_dec v__7 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in (let (imm7 :: 7 bits) = ((subrange_vec_dec v__7 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in (let (imm5 :: 5 bits) = ((subrange_vec_dec v__7 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec imm7 ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec imm5 ( 0b0100111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | STORE_FP ((v__8, rs2, rs1, DOUBLE)) => is_RV32D_or_RV64D () \ ((\ (w__57 :: bool) . if w__57 then (let (imm7 :: 7 bits) = ((subrange_vec_dec v__8 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in (let (imm7 :: 7 bits) = ((subrange_vec_dec v__8 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in (let (imm5 :: 5 bits) = ((subrange_vec_dec v__8 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec imm7 ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec imm5 ( 0b0100111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, FMADD_S)) => is_RV32F_or_RV64F () \ ((\ (w__60 :: bool) . if w__60 then return ((concat_vec rs3 ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1000011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 27 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, FMSUB_S)) => is_RV32F_or_RV64F () \ ((\ (w__63 :: bool) . if w__63 then return ((concat_vec rs3 ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1000111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 27 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, FNMSUB_S)) => is_RV32F_or_RV64F () \ ((\ (w__66 :: bool) . if w__66 then return ((concat_vec rs3 ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1001011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 27 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, FNMADD_S)) => is_RV32F_or_RV64F () \ ((\ (w__69 :: bool) . if w__69 then return ((concat_vec rs3 ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1001111 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 27 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, FADD_S)) => is_RV32F_or_RV64F () \ ((\ (w__72 :: bool) . if w__72 then return ((concat_vec ( 0b0000000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, FSUB_S)) => is_RV32F_or_RV64F () \ ((\ (w__75 :: bool) . if w__75 then return ((concat_vec ( 0b0000100 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, FMUL_S)) => is_RV32F_or_RV64F () \ ((\ (w__78 :: bool) . if w__78 then return ((concat_vec ( 0b0001000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, FDIV_S)) => is_RV32F_or_RV64F () \ ((\ (w__81 :: bool) . if w__81 then return ((concat_vec ( 0b0001100 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FSQRT_S)) => is_RV32F_or_RV64F () \ ((\ (w__84 :: bool) . if w__84 then return ((concat_vec ( 0b0101100 :: 7 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_W_S)) => is_RV32F_or_RV64F () \ ((\ (w__87 :: bool) . if w__87 then return ((concat_vec ( 0b1100000 :: 7 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_WU_S)) => is_RV32F_or_RV64F () \ ((\ (w__90 :: bool) . if w__90 then return ((concat_vec ( 0b1100000 :: 7 Word.word) ((concat_vec ( 0b00001 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_W)) => is_RV32F_or_RV64F () \ ((\ (w__93 :: bool) . if w__93 then return ((concat_vec ( 0b1101000 :: 7 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_WU)) => is_RV32F_or_RV64F () \ ((\ (w__96 :: bool) . if w__96 then return ((concat_vec ( 0b1101000 :: 7 Word.word) ((concat_vec ( 0b00001 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_L_S)) => is_RV64F () \ ((\ (w__99 :: bool) . if w__99 then return ((concat_vec ( 0b1100000 :: 7 Word.word) ((concat_vec ( 0b00010 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_LU_S)) => is_RV64F () \ ((\ (w__102 :: bool) . if w__102 then return ((concat_vec ( 0b1100000 :: 7 Word.word) ((concat_vec ( 0b00011 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_L)) => is_RV64F () \ ((\ (w__105 :: bool) . if w__105 then return ((concat_vec ( 0b1101000 :: 7 Word.word) ((concat_vec ( 0b00010 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_LU)) => is_RV64F () \ ((\ (w__108 :: bool) . if w__108 then return ((concat_vec ( 0b1101000 :: 7 Word.word) ((concat_vec ( 0b00011 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ((encdec_rounding_mode_forwards rm :: 3 Word.word)) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJ_S)) => is_RV32F_or_RV64F () \ ((\ (w__111 :: bool) . if w__111 then return ((concat_vec ( 0b0010000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJN_S)) => is_RV32F_or_RV64F () \ ((\ (w__114 :: bool) . if w__114 then return ((concat_vec ( 0b0010000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJX_S)) => is_RV32F_or_RV64F () \ ((\ (w__117 :: bool) . if w__117 then return ((concat_vec ( 0b0010000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_TYPE_S ((rs2, rs1, rd, FMIN_S)) => is_RV32F_or_RV64F () \ ((\ (w__120 :: bool) . if w__120 then return ((concat_vec ( 0b0010100 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_TYPE_S ((rs2, rs1, rd, FMAX_S)) => is_RV32F_or_RV64F () \ ((\ (w__123 :: bool) . if w__123 then return ((concat_vec ( 0b0010100 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_TYPE_S ((rs2, rs1, rd, FEQ_S)) => is_RV32F_or_RV64F () \ ((\ (w__126 :: bool) . if w__126 then return ((concat_vec ( 0b1010000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_TYPE_S ((rs2, rs1, rd, FLT_S)) => is_RV32F_or_RV64F () \ ((\ (w__129 :: bool) . if w__129 then return ((concat_vec ( 0b1010000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_BIN_TYPE_S ((rs2, rs1, rd, FLE_S)) => is_RV32F_or_RV64F () \ ((\ (w__132 :: bool) . if w__132 then return ((concat_vec ( 0b1010000 :: 7 Word.word) ((concat_vec rs2 ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_TYPE_S ((rs1, rd, FCLASS_S)) => haveFExt () \ ((\ (w__135 :: bool) . if w__135 then return ((concat_vec ( 0b1110000 :: 7 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_TYPE_S ((rs1, rd, FMV_X_W)) => haveFExt () \ ((\ (w__138 :: bool) . if w__138 then return ((concat_vec ( 0b1110000 :: 7 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | F_UN_TYPE_S ((rs1, rd, FMV_W_X)) => haveFExt () \ ((\ (w__141 :: bool) . if w__141 then return ((concat_vec ( 0b1111000 :: 7 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs1 ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec rd ( 0b1010011 :: 7 Word.word) :: 12 Word.word)) :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | ILLEGAL (s) => return s | _ => assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )))\ for ast :: " ast " definition encdec_backwards :: \(32)Word.word \((register_value),(ast),(exception))monad \ where \ encdec_backwards arg1 = ( (let v__9 = arg1 in if ((let (mappingpatterns_400 :: 7 Word.word) = ((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in encdec_uop_backwards_matches mappingpatterns_400)) then (let (imm :: 20 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 12 :: int)::ii) :: 20 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 20 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 12 :: int)::ii) :: 20 Word.word)) in (let (mappingpatterns_400 :: 7 Word.word) = ((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in encdec_uop_backwards mappingpatterns_400 \ ((\ op1 . return (UTYPE (imm, rd, op1)))))))) else if (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1101111 :: 7 Word.word)))) then (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm_8 :: 1 bits) = ((subrange_vec_dec v__9 (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word)) in (let (imm_7_0 :: 8 bits) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 12 :: int)::ii) :: 8 Word.word)) in (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in (let (imm_18_13 :: 6 bits) = ((subrange_vec_dec v__9 (( 30 :: int)::ii) (( 25 :: int)::ii) :: 6 Word.word)) in (let (imm_12_9 :: 4 bits) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 21 :: int)::ii) :: 4 Word.word)) in return (RISCV_JAL ((concat_vec imm_19 ((concat_vec imm_7_0 ((concat_vec imm_8 ((concat_vec imm_18_13 ((concat_vec imm_12_9 ( 0b0 :: 1 Word.word) :: 5 Word.word)) :: 11 Word.word)) :: 12 Word.word)) :: 20 Word.word)) :: 21 Word.word), rd))))))))) else if ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1100111 :: 7 Word.word))))))) then (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in return (RISCV_JALR (imm, rs1, rd)))))) else if (((((let (mappingpatterns_410 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_bop_backwards_matches mappingpatterns_410)) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1100011 :: 7 Word.word))))))) then (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in (let (imm7_5_0 :: 6 bits) = ((subrange_vec_dec v__9 (( 30 :: int)::ii) (( 25 :: int)::ii) :: 6 Word.word)) in (let (imm5_4_1 :: 4 bits) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 8 :: int)::ii) :: 4 Word.word)) in (let (imm5_0 :: 1 bits) = ((subrange_vec_dec v__9 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_410 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_bop_backwards mappingpatterns_410 \ ((\ op1 . return (BTYPE ((concat_vec imm7_6 ((concat_vec imm5_0 ((concat_vec imm7_5_0 ((concat_vec imm5_4_1 ( 0b0 :: 1 Word.word) :: 5 Word.word)) :: 11 Word.word)) :: 12 Word.word)) :: 13 Word.word), rs2, rs1, op1)))))))))))) else if (((((let (mappingpatterns_420 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_iop_backwards_matches mappingpatterns_420)) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0010011 :: 7 Word.word))))))) then (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (mappingpatterns_420 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_iop_backwards mappingpatterns_420 \ ((\ op1 . return (ITYPE (imm, rs1, rd, op1))))))))) else if (((((let (shamt :: 6 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((((access_vec_dec shamt (( 5 :: int)::ii))) = B0)))))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = ( 0b000000 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0010011 :: 7 Word.word))))))))))))) then (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTIOP (shamt, rs1, rd, RISCV_SLLI))))) else if (((((let (shamt :: 6 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((((access_vec_dec shamt (( 5 :: int)::ii))) = B0)))))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = ( 0b000000 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0010011 :: 7 Word.word))))))))))))) then (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTIOP (shamt, rs1, rd, RISCV_SRLI))))) else if (((((let (shamt :: 6 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((((access_vec_dec shamt (( 5 :: int)::ii))) = B0)))))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = ( 0b010000 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0010011 :: 7 Word.word))))))))))))) then (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTIOP (shamt, rs1, rd, RISCV_SRAI))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_ADD))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_SLT))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_SLTU))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_AND))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b110 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_OR))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b100 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_XOR))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_SLL))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_SRL))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_SUB))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPE (rs2, rs1, rd, RISCV_SRA))))) else and_boolM ((let (mappingpatterns_440 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_430 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_440)) then size_bits_backwards mappingpatterns_440 \ ((\ size1 . if ((bool_bits_backwards_matches mappingpatterns_430)) then bool_bits_backwards mappingpatterns_430 \ ((\ is_unsigned . return (((((((word_width_bytes size1)) < (( 4 :: int)::ii))) \ (((((\ is_unsigned)) \ ((((word_width_bytes size1)) \ (( 4 :: int)::ii))))))))))) else return False)) else return False))) (return (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000011 :: 7 Word.word))))) \ ((\ (w__2 :: bool) . if w__2 then (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (mappingpatterns_440 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_430 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in size_bits_backwards mappingpatterns_440 \ ((\ size1 . bool_bits_backwards mappingpatterns_430 \ ((\ is_unsigned . return (LOAD (imm, rs1, rd, is_unsigned, size1, False, False)))))))))))) else and_boolM ((let (mappingpatterns_450 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_450)) then size_bits_backwards mappingpatterns_450 \ ((\ size1 . return ((((word_width_bytes size1)) \ (( 4 :: int)::ii))))) else return False)) (return ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0100011 :: 7 Word.word)))))))) \ ((\ (w__4 :: bool) . if w__4 then (let (imm7 :: 7 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (imm7 :: 7 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in (let (imm5 :: 5 bits) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_450 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in size_bits_backwards mappingpatterns_450 \ ((\ size1 . return (STORE ((concat_vec imm7 imm5 :: 12 Word.word), rs2, rs1, size1, False, False)))))))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word)))))))))) then (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in return (ADDIW (imm, rs1, rd)))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTW (shamt, rs1, rd, RISCV_SLLI))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTW (shamt, rs1, rd, RISCV_SRLI))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTW (shamt, rs1, rd, RISCV_SRAI))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPEW (rs2, rs1, rd, RISCV_ADDW))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPEW (rs2, rs1, rd, RISCV_SUBW))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPEW (rs2, rs1, rd, RISCV_SLLW))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPEW (rs2, rs1, rd, RISCV_SRLW))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (RTYPEW (rs2, rs1, rd, RISCV_SRAW))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTIWOP (shamt, rs1, rd, RISCV_SLLIW))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTIWOP (shamt, rs1, rd, RISCV_SRLIW))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (SHIFTIWOP (shamt, rs1, rd, RISCV_SRAIW))))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) = ( 0x0 :: 4 Word.word)))) \ (((((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 0 :: int)::ii) :: 20 Word.word)) = ( 0x0000F :: 20 Word.word))))))) then (let (succ :: 4 Word.word) = ((subrange_vec_dec v__9 (( 23 :: int)::ii) (( 20 :: int)::ii) :: 4 Word.word)) in (let (pred :: 4 Word.word) = ((subrange_vec_dec v__9 (( 27 :: int)::ii) (( 24 :: int)::ii) :: 4 Word.word)) in return (FENCE (pred, succ)))) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) = ( 0x8 :: 4 Word.word)))) \ (((((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 0 :: int)::ii) :: 20 Word.word)) = ( 0x0000F :: 20 Word.word))))))) then (let (succ :: 4 Word.word) = ((subrange_vec_dec v__9 (( 23 :: int)::ii) (( 20 :: int)::ii) :: 4 Word.word)) in (let (pred :: 4 Word.word) = ((subrange_vec_dec v__9 (( 27 :: int)::ii) (( 24 :: int)::ii) :: 4 Word.word)) in return (FENCE_TSO (pred, succ)))) else if (((v__9 = ( 0x0000100F :: 32 Word.word)))) then return (FENCEI () ) else if (((v__9 = ( 0x00000073 :: 32 Word.word)))) then return (ECALL () ) else if (((v__9 = ( 0x30200073 :: 32 Word.word)))) then return (MRET () ) else if (((v__9 = ( 0x10200073 :: 32 Word.word)))) then return (SRET () ) else if (((v__9 = ( 0x00100073 :: 32 Word.word)))) then return (EBREAK () ) else if (((v__9 = ( 0x10500073 :: 32 Word.word)))) then return (WFI () ) else if ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0001001 :: 7 Word.word)))) \ (((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 0 :: int)::ii) :: 15 Word.word)) = ( 0b000000001110011 :: 15 Word.word))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in return (SFENCE_VMA (rs1, rs2)))) else and_boolM ((let (mappingpatterns_480 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_470 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_460 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_480)) then size_bits_backwards mappingpatterns_480 \ ((\ size1 . if ((bool_bits_backwards_matches mappingpatterns_470)) then bool_bits_backwards mappingpatterns_470 \ ((\ rl . if ((bool_bits_backwards_matches mappingpatterns_460)) then bool_bits_backwards mappingpatterns_460 \ ((\ aq . return ((amo_width_valid size1)))) else return False)) else return False)) else return False)))) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) = ( 0b00010 :: 5 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) = ( 0b00000 :: 5 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0101111 :: 7 Word.word)))))))))))))) \ ((\ (w__8 :: bool) . if w__8 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_480 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_470 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_460 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in size_bits_backwards mappingpatterns_480 \ ((\ size1 . bool_bits_backwards mappingpatterns_470 \ ((\ rl . bool_bits_backwards mappingpatterns_460 \ ((\ aq . return (LOADRES (aq, rl, rs1, size1, rd))))))))))))) else and_boolM ((let (mappingpatterns_510 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_500 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_490 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_510)) then size_bits_backwards mappingpatterns_510 \ ((\ size1 . if ((bool_bits_backwards_matches mappingpatterns_500)) then bool_bits_backwards mappingpatterns_500 \ ((\ rl . if ((bool_bits_backwards_matches mappingpatterns_490)) then bool_bits_backwards mappingpatterns_490 \ ((\ aq . return ((amo_width_valid size1)))) else return False)) else return False)) else return False)))) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) = ( 0b00011 :: 5 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0101111 :: 7 Word.word))))))))))) \ ((\ (w__12 :: bool) . if w__12 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_510 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_500 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_490 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in size_bits_backwards mappingpatterns_510 \ ((\ size1 . bool_bits_backwards mappingpatterns_500 \ ((\ rl . bool_bits_backwards mappingpatterns_490 \ ((\ aq . return (STORECON (aq, rl, rs2, rs1, size1, rd)))))))))))))) else and_boolM ((let (mappingpatterns_520 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_550 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_540 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_530 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_520 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_550)) then size_bits_backwards mappingpatterns_550 \ ((\ size1 . if ((bool_bits_backwards_matches mappingpatterns_540)) then bool_bits_backwards mappingpatterns_540 \ ((\ rl . if ((bool_bits_backwards_matches mappingpatterns_530)) then bool_bits_backwards mappingpatterns_530 \ ((\ aq . if ((encdec_amoop_backwards_matches mappingpatterns_520)) then encdec_amoop_backwards mappingpatterns_520 \ ((\ op1 . return ((amo_width_valid size1)))) else return False)) else return False)) else return False)) else return False)))))) (return ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0101111 :: 7 Word.word)))))))) \ ((\ (w__17 :: bool) . if w__17 then (let (mappingpatterns_520 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_550 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_540 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_530 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_520 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in size_bits_backwards mappingpatterns_550 \ ((\ size1 . bool_bits_backwards mappingpatterns_540 \ ((\ rl . bool_bits_backwards mappingpatterns_530 \ ((\ aq . encdec_amoop_backwards mappingpatterns_520 \ ((\ op1 . return (AMO (op1, aq, rl, rs2, rs1, size1, rd)))))))))))))))))) else if (((((let (mappingpatterns_560 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_mul_op_backwards_matches mappingpatterns_560)) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_560 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_mul_op_backwards mappingpatterns_560 \ ((\ varstup . (let (high, signed1, signed2) = varstup in return (MUL (rs2, rs1, rd, high, signed1, signed2))))))))) else if (((((let (mappingpatterns_570 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards_matches mappingpatterns_570)) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word))))))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_570 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards mappingpatterns_570 \ ((\ s . return (DIV (rs2, rs1, rd, s)))))))) else if (((((let (mappingpatterns_580 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards_matches mappingpatterns_580)) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word))))))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_580 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards mappingpatterns_580 \ ((\ s . return (REM (rs2, rs1, rd, s)))))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (MULW (rs2, rs1, rd))))) else and_boolM ((let (mappingpatterns_590 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in if ((bool_not_bits_backwards_matches mappingpatterns_590)) then bool_not_bits_backwards mappingpatterns_590 \ ((\ s . return ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))))) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))) \ ((\ (w__19 :: bool) . if w__19 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_590 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards mappingpatterns_590 \ ((\ s . return (DIVW (rs2, rs1, rd, s)))))))) else and_boolM ((let (mappingpatterns_600 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in if ((bool_not_bits_backwards_matches mappingpatterns_600)) then bool_not_bits_backwards mappingpatterns_600 \ ((\ s . return ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))))) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))) \ ((\ (w__21 :: bool) . if w__21 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_600 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards mappingpatterns_600 \ ((\ s . return (REMW (rs2, rs1, rd, s)))))))) else and_boolM ((let (mappingpatterns_620 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_610 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in if ((encdec_csrop_backwards_matches mappingpatterns_620)) then encdec_csrop_backwards mappingpatterns_620 \ ((\ op1 . return ((bool_bits_backwards_matches mappingpatterns_610)))) else return False))) (return (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1110011 :: 7 Word.word))))) \ ((\ (w__23 :: bool) . if w__23 then (let (csr :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (csr :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (mappingpatterns_620 :: 2 Word.word) = ((subrange_vec_dec v__9 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_610 :: 1 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in encdec_csrop_backwards mappingpatterns_620 \ ((\ op1 . bool_bits_backwards mappingpatterns_610 \ ((\ is_imm . return (CSR (csr, rs1, rd, is_imm, op1)))))))))))) else if (((v__9 = ( 0x00200073 :: 32 Word.word)))) then return (URET () ) else if (((((let (fm :: 4 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) in (let (rs :: regidx) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (fm :: 4 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) in (((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word))))))))))))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0001111 :: 7 Word.word)))))))))) then (let (fm :: 4 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) in (let (succ :: 4 bits) = ((subrange_vec_dec v__9 (( 23 :: int)::ii) (( 20 :: int)::ii) :: 4 Word.word)) in (let (rs :: regidx) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (pred :: 4 bits) = ((subrange_vec_dec v__9 (( 27 :: int)::ii) (( 24 :: int)::ii) :: 4 Word.word)) in (let (fm :: 4 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) in return (FENCE_RESERVED (fm, pred, succ, rs, rd)))))))) else if (((((let (imm :: 12 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs :: regidx) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in ((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg)))))))))))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0001111 :: 7 Word.word)))))))))) then (let (imm :: 12 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs :: regidx) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in return (FENCEI_RESERVED (imm, rs, rd)))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000111 :: 7 Word.word)))))))) \ ((\ (w__25 :: bool) . if w__25 then (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in return (LOAD_FP (imm, rs1, rd, WORD)))))) else and_boolM ((is_RV32D_or_RV64D () )) (return ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000111 :: 7 Word.word)))))))) \ ((\ (w__27 :: bool) . if w__27 then (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in return (LOAD_FP (imm, rs1, rd, DOUBLE)))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0100111 :: 7 Word.word)))))))) \ ((\ (w__29 :: bool) . if w__29 then (let (imm7 :: 7 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (imm7 :: 7 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in (let (imm5 :: 5 bits) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (STORE_FP ((concat_vec imm7 imm5 :: 12 Word.word), rs2, rs1, WORD))))))) else and_boolM ((is_RV32D_or_RV64D () )) (return ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0100111 :: 7 Word.word)))))))) \ ((\ (w__31 :: bool) . if w__31 then (let (imm7 :: 7 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (imm7 :: 7 bits) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in (let (imm5 :: 5 bits) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (STORE_FP ((concat_vec imm7 imm5 :: 12 Word.word), rs2, rs1, DOUBLE))))))) else and_boolM ((let (mappingpatterns_630 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_630)) then encdec_rounding_mode_backwards mappingpatterns_630 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 25 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1000011 :: 7 Word.word)))))))) \ ((\ (w__34 :: bool) . if w__34 then (let (rs3 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs3 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_630 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_630 \ ((\ rm . return (F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FMADD_S)))))))))) else and_boolM ((let (mappingpatterns_640 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_640)) then encdec_rounding_mode_backwards mappingpatterns_640 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 25 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1000111 :: 7 Word.word)))))))) \ ((\ (w__37 :: bool) . if w__37 then (let (rs3 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs3 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_640 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_640 \ ((\ rm . return (F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FMSUB_S)))))))))) else and_boolM ((let (mappingpatterns_650 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_650)) then encdec_rounding_mode_backwards mappingpatterns_650 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 25 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1001011 :: 7 Word.word)))))))) \ ((\ (w__40 :: bool) . if w__40 then (let (rs3 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs3 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_650 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_650 \ ((\ rm . return (F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FNMSUB_S)))))))))) else and_boolM ((let (mappingpatterns_660 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_660)) then encdec_rounding_mode_backwards mappingpatterns_660 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 26 :: int)::ii) (( 25 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1001111 :: 7 Word.word)))))))) \ ((\ (w__43 :: bool) . if w__43 then (let (rs3 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs3 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_660 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_660 \ ((\ rm . return (F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, FNMADD_S)))))))))) else and_boolM ((let (mappingpatterns_670 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_670)) then encdec_rounding_mode_backwards mappingpatterns_670 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__46 :: bool) . if w__46 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_670 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_670 \ ((\ rm . return (F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FADD_S)))))))) else and_boolM ((let (mappingpatterns_680 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_680)) then encdec_rounding_mode_backwards mappingpatterns_680 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000100 :: 7 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__49 :: bool) . if w__49 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_680 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_680 \ ((\ rm . return (F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FSUB_S)))))))) else and_boolM ((let (mappingpatterns_690 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_690)) then encdec_rounding_mode_backwards mappingpatterns_690 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0001000 :: 7 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__52 :: bool) . if w__52 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_690 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_690 \ ((\ rm . return (F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FMUL_S)))))))) else and_boolM ((let (mappingpatterns_700 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_700)) then encdec_rounding_mode_backwards mappingpatterns_700 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0001100 :: 7 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__55 :: bool) . if w__55 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_700 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_700 \ ((\ rm . return (F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, FDIV_S)))))))) else and_boolM ((let (mappingpatterns_710 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_710)) then encdec_rounding_mode_backwards mappingpatterns_710 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0x580 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__58 :: bool) . if w__58 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_710 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_710 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FSQRT_S))))))) else and_boolM ((let (mappingpatterns_720 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_720)) then encdec_rounding_mode_backwards mappingpatterns_720 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xC00 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__61 :: bool) . if w__61 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_720 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_720 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_W_S))))))) else and_boolM ((let (mappingpatterns_730 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_730)) then encdec_rounding_mode_backwards mappingpatterns_730 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xC01 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__64 :: bool) . if w__64 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_730 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_730 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_WU_S))))))) else and_boolM ((let (mappingpatterns_740 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_740)) then encdec_rounding_mode_backwards mappingpatterns_740 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xD00 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__67 :: bool) . if w__67 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_740 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_740 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_W))))))) else and_boolM ((let (mappingpatterns_750 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_750)) then encdec_rounding_mode_backwards mappingpatterns_750 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xD01 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__70 :: bool) . if w__70 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_750 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_750 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_WU))))))) else and_boolM ((let (mappingpatterns_760 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_760)) then encdec_rounding_mode_backwards mappingpatterns_760 \ ((\ rm . is_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xC02 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__73 :: bool) . if w__73 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_760 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_760 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_L_S))))))) else and_boolM ((let (mappingpatterns_770 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_770)) then encdec_rounding_mode_backwards mappingpatterns_770 \ ((\ rm . is_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xC03 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__76 :: bool) . if w__76 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_770 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_770 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_LU_S))))))) else and_boolM ((let (mappingpatterns_780 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_780)) then encdec_rounding_mode_backwards mappingpatterns_780 \ ((\ rm . is_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xD02 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__79 :: bool) . if w__79 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_780 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_780 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_L))))))) else and_boolM ((let (mappingpatterns_790 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_790)) then encdec_rounding_mode_backwards mappingpatterns_790 \ ((\ rm . is_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xD03 :: 12 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__82 :: bool) . if w__82 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_790 :: 3 Word.word) = ((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_790 \ ((\ rm . return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_LU))))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__84 :: bool) . if w__84 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJ_S))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__86 :: bool) . if w__86 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJN_S))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__88 :: bool) . if w__88 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJX_S))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010100 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__90 :: bool) . if w__90 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_BIN_TYPE_S (rs2, rs1, rd, FMIN_S))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010100 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__92 :: bool) . if w__92 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_BIN_TYPE_S (rs2, rs1, rd, FMAX_S))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b1010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__94 :: bool) . if w__94 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_BIN_TYPE_S (rs2, rs1, rd, FEQ_S))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b1010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__96 :: bool) . if w__96 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_BIN_TYPE_S (rs2, rs1, rd, FLT_S))))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b1010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__98 :: bool) . if w__98 then (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_BIN_TYPE_S (rs2, rs1, rd, FLE_S))))) else and_boolM ((haveFExt () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xE00 :: 12 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__100 :: bool) . if w__100 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_UN_TYPE_S (rs1, rd, FCLASS_S)))) else and_boolM ((haveFExt () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xE00 :: 12 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__102 :: bool) . if w__102 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (F_UN_TYPE_S (rs1, rd, FMV_X_W)))) else and_boolM ((haveFExt () )) (return ((((((((subrange_vec_dec v__9 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xF00 :: 12 Word.word)))) \ ((((((((subrange_vec_dec v__9 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__9 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__104 :: bool) . return (if w__104 then (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__9 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: 5 Word.word) = ((subrange_vec_dec v__9 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in F_UN_TYPE_S (rs1, rd, FMV_W_X))) else ILLEGAL v__9)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))\ for arg1 :: "(32)Word.word " definition encdec_forwards_matches :: \ ast \((register_value),(bool),(exception))monad \ where \ encdec_forwards_matches ast = ( (let arg1 = ast in (case arg1 of UTYPE ((imm, rd, op1)) => return True | RISCV_JAL ((v__347, rd)) => return (if (((((subrange_vec_dec v__347 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then True else False) | RISCV_JALR ((imm, rs1, rd)) => return True | BTYPE ((v__349, rs2, rs1, op1)) => return (if (((((subrange_vec_dec v__349 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then True else False) | ITYPE ((imm, rs1, rd, op1)) => return True | SHIFTIOP ((shamt, rs1, rd, RISCV_SLLI)) => return True | SHIFTIOP ((shamt, rs1, rd, RISCV_SRLI)) => return True | SHIFTIOP ((shamt, rs1, rd, RISCV_SRAI)) => return True | RTYPE ((rs2, rs1, rd, RISCV_ADD)) => return True | RTYPE ((rs2, rs1, rd, RISCV_SLT)) => return True | RTYPE ((rs2, rs1, rd, RISCV_SLTU)) => return True | RTYPE ((rs2, rs1, rd, RISCV_AND)) => return True | RTYPE ((rs2, rs1, rd, RISCV_OR)) => return True | RTYPE ((rs2, rs1, rd, RISCV_XOR)) => return True | RTYPE ((rs2, rs1, rd, RISCV_SLL)) => return True | RTYPE ((rs2, rs1, rd, RISCV_SRL)) => return True | RTYPE ((rs2, rs1, rd, RISCV_SUB)) => return True | RTYPE ((rs2, rs1, rd, RISCV_SRA)) => return True | LOAD ((imm, rs1, rd, is_unsigned, size1, False, False)) => return (if (((((((word_width_bytes size1)) < (( 4 :: int)::ii))) \ (((((\ is_unsigned)) \ ((((word_width_bytes size1)) \ (( 4 :: int)::ii))))))))) then True else False) | STORE ((v__351, rs2, rs1, size1, False, False)) => return (if ((((word_width_bytes size1)) \ (( 4 :: int)::ii))) then True else False) | ADDIW ((imm, rs1, rd)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | SHIFTW ((shamt, rs1, rd, RISCV_SLLI)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | SHIFTW ((shamt, rs1, rd, RISCV_SRLI)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | SHIFTW ((shamt, rs1, rd, RISCV_SRAI)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | RTYPEW ((rs2, rs1, rd, RISCV_ADDW)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | RTYPEW ((rs2, rs1, rd, RISCV_SUBW)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | RTYPEW ((rs2, rs1, rd, RISCV_SLLW)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | RTYPEW ((rs2, rs1, rd, RISCV_SRLW)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | RTYPEW ((rs2, rs1, rd, RISCV_SRAW)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | SHIFTIWOP ((shamt, rs1, rd, RISCV_SLLIW)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | SHIFTIWOP ((shamt, rs1, rd, RISCV_SRLIW)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | SHIFTIWOP ((shamt, rs1, rd, RISCV_SRAIW)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | FENCE ((pred, succ)) => return True | FENCE_TSO ((pred, succ)) => return True | FENCEI (_) => return True | ECALL (_) => return True | MRET (_) => return True | SRET (_) => return True | EBREAK (_) => return True | WFI (_) => return True | SFENCE_VMA ((rs1, rs2)) => return True | LOADRES ((aq, rl, rs1, size1, rd)) => return (if ((amo_width_valid size1)) then True else False) | STORECON ((aq, rl, rs2, rs1, size1, rd)) => return (if ((amo_width_valid size1)) then True else False) | AMO ((op1, aq, rl, rs2, rs1, size1, rd)) => return (if ((amo_width_valid size1)) then True else False) | MUL ((rs2, rs1, rd, high, signed1, signed2)) => return True | DIV ((rs2, rs1, rd, s)) => return True | REM ((rs2, rs1, rd, s)) => return True | MULW ((rs2, rs1, rd)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | DIVW ((rs2, rs1, rd, s)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | REMW ((rs2, rs1, rd, s)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | CSR ((csr, rs1, rd, is_imm, op1)) => return True | URET (_) => return True | FENCE_RESERVED ((fm, pred, succ, rs, rd)) => return (if (((((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word)))))))))) then True else False) | FENCEI_RESERVED ((imm, rs, rd)) => return (if ((((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg))))))))) then True else False) | LOAD_FP ((imm, rs1, rd, WORD)) => is_RV32F_or_RV64F () \ ((\ (w__0 :: bool) . return (if w__0 then True else False))) | LOAD_FP ((imm, rs1, rd, DOUBLE)) => is_RV32D_or_RV64D () \ ((\ (w__1 :: bool) . return (if w__1 then True else False))) | STORE_FP ((v__352, rs2, rs1, WORD)) => is_RV32F_or_RV64F () \ ((\ (w__2 :: bool) . return (if w__2 then True else False))) | STORE_FP ((v__353, rs2, rs1, DOUBLE)) => is_RV32D_or_RV64D () \ ((\ (w__3 :: bool) . return (if w__3 then True else False))) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, FMADD_S)) => is_RV32F_or_RV64F () \ ((\ (w__4 :: bool) . return (if w__4 then True else False))) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, FMSUB_S)) => is_RV32F_or_RV64F () \ ((\ (w__5 :: bool) . return (if w__5 then True else False))) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, FNMSUB_S)) => is_RV32F_or_RV64F () \ ((\ (w__6 :: bool) . return (if w__6 then True else False))) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, FNMADD_S)) => is_RV32F_or_RV64F () \ ((\ (w__7 :: bool) . return (if w__7 then True else False))) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, FADD_S)) => is_RV32F_or_RV64F () \ ((\ (w__8 :: bool) . return (if w__8 then True else False))) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, FSUB_S)) => is_RV32F_or_RV64F () \ ((\ (w__9 :: bool) . return (if w__9 then True else False))) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, FMUL_S)) => is_RV32F_or_RV64F () \ ((\ (w__10 :: bool) . return (if w__10 then True else False))) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, FDIV_S)) => is_RV32F_or_RV64F () \ ((\ (w__11 :: bool) . return (if w__11 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FSQRT_S)) => is_RV32F_or_RV64F () \ ((\ (w__12 :: bool) . return (if w__12 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_W_S)) => is_RV32F_or_RV64F () \ ((\ (w__13 :: bool) . return (if w__13 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_WU_S)) => is_RV32F_or_RV64F () \ ((\ (w__14 :: bool) . return (if w__14 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_W)) => is_RV32F_or_RV64F () \ ((\ (w__15 :: bool) . return (if w__15 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_WU)) => is_RV32F_or_RV64F () \ ((\ (w__16 :: bool) . return (if w__16 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_L_S)) => is_RV64F () \ ((\ (w__17 :: bool) . return (if w__17 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_LU_S)) => is_RV64F () \ ((\ (w__18 :: bool) . return (if w__18 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_L)) => is_RV64F () \ ((\ (w__19 :: bool) . return (if w__19 then True else False))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_LU)) => is_RV64F () \ ((\ (w__20 :: bool) . return (if w__20 then True else False))) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJ_S)) => is_RV32F_or_RV64F () \ ((\ (w__21 :: bool) . return (if w__21 then True else False))) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJN_S)) => is_RV32F_or_RV64F () \ ((\ (w__22 :: bool) . return (if w__22 then True else False))) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJX_S)) => is_RV32F_or_RV64F () \ ((\ (w__23 :: bool) . return (if w__23 then True else False))) | F_BIN_TYPE_S ((rs2, rs1, rd, FMIN_S)) => is_RV32F_or_RV64F () \ ((\ (w__24 :: bool) . return (if w__24 then True else False))) | F_BIN_TYPE_S ((rs2, rs1, rd, FMAX_S)) => is_RV32F_or_RV64F () \ ((\ (w__25 :: bool) . return (if w__25 then True else False))) | F_BIN_TYPE_S ((rs2, rs1, rd, FEQ_S)) => is_RV32F_or_RV64F () \ ((\ (w__26 :: bool) . return (if w__26 then True else False))) | F_BIN_TYPE_S ((rs2, rs1, rd, FLT_S)) => is_RV32F_or_RV64F () \ ((\ (w__27 :: bool) . return (if w__27 then True else False))) | F_BIN_TYPE_S ((rs2, rs1, rd, FLE_S)) => is_RV32F_or_RV64F () \ ((\ (w__28 :: bool) . return (if w__28 then True else False))) | F_UN_TYPE_S ((rs1, rd, FCLASS_S)) => haveFExt () \ ((\ (w__29 :: bool) . return (if w__29 then True else False))) | F_UN_TYPE_S ((rs1, rd, FMV_X_W)) => haveFExt () \ ((\ (w__30 :: bool) . return (if w__30 then True else False))) | F_UN_TYPE_S ((rs1, rd, FMV_W_X)) => haveFExt () \ ((\ (w__31 :: bool) . return (if w__31 then True else False))) | ILLEGAL (s) => return True | _ => return False )))\ for ast :: " ast " definition encdec_backwards_matches :: \(32)Word.word \((register_value),(bool),(exception))monad \ where \ encdec_backwards_matches arg1 = ( (let v__354 = arg1 in if ((let (mappingpatterns_00 :: 7 Word.word) = ((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in encdec_uop_backwards_matches mappingpatterns_00)) then (let (mappingpatterns_00 :: 7 Word.word) = ((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in encdec_uop_backwards mappingpatterns_00 \ ((\ op1 . return True))) else if (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1101111 :: 7 Word.word)))) then return True else if ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1100111 :: 7 Word.word))))))) then return True else if (((((let (mappingpatterns_10 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_bop_backwards_matches mappingpatterns_10)) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1100011 :: 7 Word.word))))))) then (let (mappingpatterns_10 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_bop_backwards mappingpatterns_10 \ ((\ op1 . return True))) else if (((((let (mappingpatterns_20 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_iop_backwards_matches mappingpatterns_20)) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0010011 :: 7 Word.word))))))) then (let (mappingpatterns_20 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_iop_backwards mappingpatterns_20 \ ((\ op1 . return True))) else if (((((let (shamt :: 6 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((((access_vec_dec shamt (( 5 :: int)::ii))) = B0)))))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = ( 0b000000 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0010011 :: 7 Word.word))))))))))))) then return True else if (((((let (shamt :: 6 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((((access_vec_dec shamt (( 5 :: int)::ii))) = B0)))))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = ( 0b000000 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0010011 :: 7 Word.word))))))))))))) then return True else if (((((let (shamt :: 6 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in (((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((((access_vec_dec shamt (( 5 :: int)::ii))) = B0)))))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = ( 0b010000 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0010011 :: 7 Word.word))))))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b110 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b100 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then return True else and_boolM ((let (mappingpatterns_40 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_30 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_40)) then size_bits_backwards mappingpatterns_40 \ ((\ size1 . if ((bool_bits_backwards_matches mappingpatterns_30)) then bool_bits_backwards mappingpatterns_30 \ ((\ is_unsigned . return (((((((word_width_bytes size1)) < (( 4 :: int)::ii))) \ (((((\ is_unsigned)) \ ((((word_width_bytes size1)) \ (( 4 :: int)::ii))))))))))) else return False)) else return False))) (return (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000011 :: 7 Word.word))))) \ ((\ (w__2 :: bool) . if w__2 then (let (mappingpatterns_40 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_30 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in size_bits_backwards mappingpatterns_40 \ ((\ size1 . bool_bits_backwards mappingpatterns_30 \ ((\ is_unsigned . return True)))))) else and_boolM ((let (mappingpatterns_50 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_50)) then size_bits_backwards mappingpatterns_50 \ ((\ size1 . return ((((word_width_bytes size1)) \ (( 4 :: int)::ii))))) else return False)) (return ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0100011 :: 7 Word.word)))))))) \ ((\ (w__4 :: bool) . if w__4 then (let (mappingpatterns_50 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in size_bits_backwards mappingpatterns_50 \ ((\ size1 . return True))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word)))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0100000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0011011 :: 7 Word.word))))))))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) = ( 0x0 :: 4 Word.word)))) \ (((((subrange_vec_dec v__354 (( 19 :: int)::ii) (( 0 :: int)::ii) :: 20 Word.word)) = ( 0x0000F :: 20 Word.word))))))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) = ( 0x8 :: 4 Word.word)))) \ (((((subrange_vec_dec v__354 (( 19 :: int)::ii) (( 0 :: int)::ii) :: 20 Word.word)) = ( 0x0000F :: 20 Word.word))))))) then return True else if (((v__354 = ( 0x0000100F :: 32 Word.word)))) then return True else if (((v__354 = ( 0x00000073 :: 32 Word.word)))) then return True else if (((v__354 = ( 0x30200073 :: 32 Word.word)))) then return True else if (((v__354 = ( 0x10200073 :: 32 Word.word)))) then return True else if (((v__354 = ( 0x00100073 :: 32 Word.word)))) then return True else if (((v__354 = ( 0x10500073 :: 32 Word.word)))) then return True else if ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0001001 :: 7 Word.word)))) \ (((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 0 :: int)::ii) :: 15 Word.word)) = ( 0b000000001110011 :: 15 Word.word))))))) then return True else and_boolM ((let (mappingpatterns_80 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_70 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_60 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_80)) then size_bits_backwards mappingpatterns_80 \ ((\ size1 . if ((bool_bits_backwards_matches mappingpatterns_70)) then bool_bits_backwards mappingpatterns_70 \ ((\ rl . if ((bool_bits_backwards_matches mappingpatterns_60)) then bool_bits_backwards mappingpatterns_60 \ ((\ aq . return ((amo_width_valid size1)))) else return False)) else return False)) else return False)))) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) = ( 0b00010 :: 5 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) = ( 0b00000 :: 5 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0101111 :: 7 Word.word)))))))))))))) \ ((\ (w__8 :: bool) . if w__8 then (let (mappingpatterns_80 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_70 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_60 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in size_bits_backwards mappingpatterns_80 \ ((\ size1 . bool_bits_backwards mappingpatterns_70 \ ((\ rl . bool_bits_backwards mappingpatterns_60 \ ((\ aq . return True))))))))) else and_boolM ((let (mappingpatterns_90 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_110 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_100 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_110)) then size_bits_backwards mappingpatterns_110 \ ((\ size1 . if ((bool_bits_backwards_matches mappingpatterns_100)) then bool_bits_backwards mappingpatterns_100 \ ((\ rl . if ((bool_bits_backwards_matches mappingpatterns_90)) then bool_bits_backwards mappingpatterns_90 \ ((\ aq . return ((amo_width_valid size1)))) else return False)) else return False)) else return False)))) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) = ( 0b00011 :: 5 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0101111 :: 7 Word.word))))))))))) \ ((\ (w__12 :: bool) . if w__12 then (let (mappingpatterns_90 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_110 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_100 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in size_bits_backwards mappingpatterns_110 \ ((\ size1 . bool_bits_backwards mappingpatterns_100 \ ((\ rl . bool_bits_backwards mappingpatterns_90 \ ((\ aq . return True))))))))) else and_boolM ((let (mappingpatterns_120 :: 5 Word.word) = ((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_150 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_140 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_130 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_120 :: 5 Word.word) = ((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in if ((size_bits_backwards_matches mappingpatterns_150)) then size_bits_backwards mappingpatterns_150 \ ((\ size1 . if ((bool_bits_backwards_matches mappingpatterns_140)) then bool_bits_backwards mappingpatterns_140 \ ((\ rl . if ((bool_bits_backwards_matches mappingpatterns_130)) then bool_bits_backwards mappingpatterns_130 \ ((\ aq . if ((encdec_amoop_backwards_matches mappingpatterns_120)) then encdec_amoop_backwards mappingpatterns_120 \ ((\ op1 . return ((amo_width_valid size1)))) else return False)) else return False)) else return False)) else return False)))))) (return ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0101111 :: 7 Word.word)))))))) \ ((\ (w__17 :: bool) . if w__17 then (let (mappingpatterns_120 :: 5 Word.word) = ((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let (mappingpatterns_150 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_140 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_130 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let (mappingpatterns_120 :: 5 Word.word) = ((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in size_bits_backwards mappingpatterns_150 \ ((\ size1 . bool_bits_backwards mappingpatterns_140 \ ((\ rl . bool_bits_backwards mappingpatterns_130 \ ((\ aq . encdec_amoop_backwards mappingpatterns_120 \ ((\ op1 . return True))))))))))))) else if (((((let (mappingpatterns_160 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_mul_op_backwards_matches mappingpatterns_160)) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word)))))))))) then (let (mappingpatterns_160 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_mul_op_backwards mappingpatterns_160 \ ((\ varstup . (let (high, signed1, signed2) = varstup in return True)))) else if (((((let (mappingpatterns_170 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards_matches mappingpatterns_170)) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word))))))))))))) then (let (mappingpatterns_170 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards mappingpatterns_170 \ ((\ s . return True))) else if (((((let (mappingpatterns_180 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards_matches mappingpatterns_180)) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0110011 :: 7 Word.word))))))))))))) then (let (mappingpatterns_180 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards mappingpatterns_180 \ ((\ s . return True))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))))) then return True else and_boolM ((let (mappingpatterns_190 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in if ((bool_not_bits_backwards_matches mappingpatterns_190)) then bool_not_bits_backwards mappingpatterns_190 \ ((\ s . return ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))))) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))) \ ((\ (w__19 :: bool) . if w__19 then (let (mappingpatterns_190 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards mappingpatterns_190 \ ((\ s . return True))) else and_boolM ((let (mappingpatterns_200 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in if ((bool_not_bits_backwards_matches mappingpatterns_200)) then bool_not_bits_backwards mappingpatterns_200 \ ((\ s . return ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))))) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0111011 :: 7 Word.word))))))))))) \ ((\ (w__21 :: bool) . if w__21 then (let (mappingpatterns_200 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in bool_not_bits_backwards mappingpatterns_200 \ ((\ s . return True))) else and_boolM ((let (mappingpatterns_220 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_210 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in if ((encdec_csrop_backwards_matches mappingpatterns_220)) then encdec_csrop_backwards mappingpatterns_220 \ ((\ op1 . return ((bool_bits_backwards_matches mappingpatterns_210)))) else return False))) (return (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1110011 :: 7 Word.word))))) \ ((\ (w__23 :: bool) . if w__23 then (let (mappingpatterns_220 :: 2 Word.word) = ((subrange_vec_dec v__354 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let (mappingpatterns_210 :: 1 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in encdec_csrop_backwards mappingpatterns_220 \ ((\ op1 . bool_bits_backwards mappingpatterns_210 \ ((\ is_imm . return True)))))) else if (((v__354 = ( 0x00200073 :: 32 Word.word)))) then return True else if (((((let (fm :: 4 bits) = ((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) in (let (rs :: regidx) = ((subrange_vec_dec v__354 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__354 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (fm :: 4 bits) = ((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) in (((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word))))))))))))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0001111 :: 7 Word.word)))))))))) then return True else if (((((let (imm :: 12 bits) = ((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in (let (rs :: regidx) = ((subrange_vec_dec v__354 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__354 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm :: 12 bits) = ((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in ((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg)))))))))))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0001111 :: 7 Word.word)))))))))) then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000111 :: 7 Word.word)))))))) \ ((\ (w__25 :: bool) . if w__25 then return True else and_boolM ((is_RV32D_or_RV64D () )) (return ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000111 :: 7 Word.word)))))))) \ ((\ (w__27 :: bool) . if w__27 then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0100111 :: 7 Word.word)))))))) \ ((\ (w__29 :: bool) . if w__29 then return True else and_boolM ((is_RV32D_or_RV64D () )) (return ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0100111 :: 7 Word.word)))))))) \ ((\ (w__31 :: bool) . if w__31 then return True else and_boolM ((let (mappingpatterns_230 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_230)) then encdec_rounding_mode_backwards mappingpatterns_230 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 25 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1000011 :: 7 Word.word)))))))) \ ((\ (w__34 :: bool) . if w__34 then (let (mappingpatterns_230 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_230 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_240 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_240)) then encdec_rounding_mode_backwards mappingpatterns_240 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 25 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1000111 :: 7 Word.word)))))))) \ ((\ (w__37 :: bool) . if w__37 then (let (mappingpatterns_240 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_240 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_250 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_250)) then encdec_rounding_mode_backwards mappingpatterns_250 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 25 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1001011 :: 7 Word.word)))))))) \ ((\ (w__40 :: bool) . if w__40 then (let (mappingpatterns_250 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_250 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_260 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_260)) then encdec_rounding_mode_backwards mappingpatterns_260 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 26 :: int)::ii) (( 25 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1001111 :: 7 Word.word)))))))) \ ((\ (w__43 :: bool) . if w__43 then (let (mappingpatterns_260 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_260 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_270 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_270)) then encdec_rounding_mode_backwards mappingpatterns_270 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000000 :: 7 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__46 :: bool) . if w__46 then (let (mappingpatterns_270 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_270 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_280 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_280)) then encdec_rounding_mode_backwards mappingpatterns_280 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0000100 :: 7 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__49 :: bool) . if w__49 then (let (mappingpatterns_280 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_280 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_290 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_290)) then encdec_rounding_mode_backwards mappingpatterns_290 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0001000 :: 7 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__52 :: bool) . if w__52 then (let (mappingpatterns_290 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_290 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_300 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_300)) then encdec_rounding_mode_backwards mappingpatterns_300 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0001100 :: 7 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__55 :: bool) . if w__55 then (let (mappingpatterns_300 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_300 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_310 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_310)) then encdec_rounding_mode_backwards mappingpatterns_310 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0x580 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__58 :: bool) . if w__58 then (let (mappingpatterns_310 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_310 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_320 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_320)) then encdec_rounding_mode_backwards mappingpatterns_320 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xC00 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__61 :: bool) . if w__61 then (let (mappingpatterns_320 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_320 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_330 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_330)) then encdec_rounding_mode_backwards mappingpatterns_330 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xC01 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__64 :: bool) . if w__64 then (let (mappingpatterns_330 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_330 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_340 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_340)) then encdec_rounding_mode_backwards mappingpatterns_340 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xD00 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__67 :: bool) . if w__67 then (let (mappingpatterns_340 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_340 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_350 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_350)) then encdec_rounding_mode_backwards mappingpatterns_350 \ ((\ rm . is_RV32F_or_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xD01 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__70 :: bool) . if w__70 then (let (mappingpatterns_350 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_350 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_360 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_360)) then encdec_rounding_mode_backwards mappingpatterns_360 \ ((\ rm . is_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xC02 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__73 :: bool) . if w__73 then (let (mappingpatterns_360 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_360 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_370 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_370)) then encdec_rounding_mode_backwards mappingpatterns_370 \ ((\ rm . is_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xC03 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__76 :: bool) . if w__76 then (let (mappingpatterns_370 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_370 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_380 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_380)) then encdec_rounding_mode_backwards mappingpatterns_380 \ ((\ rm . is_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xD02 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__79 :: bool) . if w__79 then (let (mappingpatterns_380 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_380 \ ((\ rm . return True))) else and_boolM ((let (mappingpatterns_390 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in if ((encdec_rounding_mode_backwards_matches mappingpatterns_390)) then encdec_rounding_mode_backwards mappingpatterns_390 \ ((\ rm . is_RV64F () )) else return False)) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xD03 :: 12 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word)))))))) \ ((\ (w__82 :: bool) . if w__82 then (let (mappingpatterns_390 :: 3 Word.word) = ((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in encdec_rounding_mode_backwards mappingpatterns_390 \ ((\ rm . return True))) else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__84 :: bool) . if w__84 then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__86 :: bool) . if w__86 then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__88 :: bool) . if w__88 then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010100 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__90 :: bool) . if w__90 then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b0010100 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__92 :: bool) . if w__92 then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b1010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__94 :: bool) . if w__94 then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b1010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__96 :: bool) . if w__96 then return True else and_boolM ((is_RV32F_or_RV64F () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = ( 0b1010000 :: 7 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__98 :: bool) . if w__98 then return True else and_boolM ((haveFExt () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xE00 :: 12 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__100 :: bool) . if w__100 then return True else and_boolM ((haveFExt () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xE00 :: 12 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__102 :: bool) . if w__102 then return True else and_boolM ((haveFExt () )) (return ((((((((subrange_vec_dec v__354 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) = ( 0xF00 :: 12 Word.word)))) \ ((((((((subrange_vec_dec v__354 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__354 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b1010011 :: 7 Word.word))))))))))) \ ((\ (w__104 :: bool) . return (if w__104 then True else True)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))\ for arg1 :: "(32)Word.word " definition encdec_compressed_forwards :: \ ast \((register_value),((16)Word.word),(exception))monad \ where \ encdec_compressed_forwards ast = ( (let arg1 = ast in (case arg1 of C_NOP (_) => return ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) | C_ADDI4SPN ((rd, v__692)) => if ((let (nz96 :: 4 bits) = ((subrange_vec_dec v__692 (( 7 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word)) in (let (nz96 :: 4 bits) = ((subrange_vec_dec v__692 (( 7 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word)) in (let (nz54 :: 2 bits) = ((subrange_vec_dec v__692 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in (let (nz3 :: 1 bits) = ((subrange_vec_dec v__692 (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word)) in (let (nz2 :: 1 bits) = ((subrange_vec_dec v__692 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in (((concat_vec nz96 ((concat_vec nz54 ((concat_vec nz3 nz2 :: 2 Word.word)) :: 4 Word.word)) :: 8 Word.word)) \ ( 0x00 :: 8 Word.word)))))))) then (let (nz96 :: 4 bits) = ((subrange_vec_dec v__692 (( 7 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word)) in (let (nz96 :: 4 bits) = ((subrange_vec_dec v__692 (( 7 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word)) in (let (nz54 :: 2 bits) = ((subrange_vec_dec v__692 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in (let (nz3 :: 1 bits) = ((subrange_vec_dec v__692 (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word)) in (let (nz2 :: 1 bits) = ((subrange_vec_dec v__692 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in return ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec nz54 ((concat_vec nz96 ((concat_vec nz2 ((concat_vec nz3 ((concat_vec rd ( 0b00 :: 2 Word.word) :: 5 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 11 Word.word)) :: 13 Word.word)) :: 16 Word.word))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LW ((v__693, rs1, rd)) => (let (ui6 :: 1 bits) = ((subrange_vec_dec v__693 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (ui6 :: 1 bits) = ((subrange_vec_dec v__693 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__693 (( 3 :: int)::ii) (( 1 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__693 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in return ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec ui53 ((concat_vec rs1 ((concat_vec ui2 ((concat_vec ui6 ((concat_vec rd ( 0b00 :: 2 Word.word) :: 5 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))) | C_LD ((v__694, rs1, rd)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__694 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (ui76 :: 2 bits) = ((subrange_vec_dec v__694 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__694 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in return ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec ui53 ((concat_vec rs1 ((concat_vec ui76 ((concat_vec rd ( 0b00 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SW ((v__695, rs1, rs2)) => (let (ui6 :: 1 bits) = ((subrange_vec_dec v__695 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (ui6 :: 1 bits) = ((subrange_vec_dec v__695 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__695 (( 3 :: int)::ii) (( 1 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__695 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in return ((concat_vec ( 0b110 :: 3 Word.word) ((concat_vec ui53 ((concat_vec rs1 ((concat_vec ui2 ((concat_vec ui6 ((concat_vec rs2 ( 0b00 :: 2 Word.word) :: 5 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))) | C_SD ((v__696, rs1, rs2)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__696 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (ui76 :: 2 bits) = ((subrange_vec_dec v__696 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__696 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in return ((concat_vec ( 0b111 :: 3 Word.word) ((concat_vec ui53 ((concat_vec rs1 ((concat_vec ui76 ((concat_vec rs2 ( 0b00 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDI ((v__697, rsd)) => if ((let (nzi5 :: 1 bits) = ((subrange_vec_dec v__697 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__697 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi40 :: 5 bits) = ((subrange_vec_dec v__697 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in ((((((concat_vec nzi5 nzi40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))))))) then (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__697 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__697 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi40 :: 5 bits) = ((subrange_vec_dec v__697 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec nzi5 ((concat_vec rsd ((concat_vec nzi40 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_JAL (v__698) => if ((((( 32 :: int)::ii) = (( 32 :: int)::ii)))) then (let (i11 :: 1 bits) = ((subrange_vec_dec v__698 (( 10 :: int)::ii) (( 10 :: int)::ii) :: 1 Word.word)) in (let (i98 :: 2 bits) = ((subrange_vec_dec v__698 (( 8 :: int)::ii) (( 7 :: int)::ii) :: 2 Word.word)) in (let (i7 :: 1 bits) = ((subrange_vec_dec v__698 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let (i6 :: 1 bits) = ((subrange_vec_dec v__698 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__698 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (i4 :: 1 bits) = ((subrange_vec_dec v__698 (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word)) in (let (i31 :: 3 bits) = ((subrange_vec_dec v__698 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in (let (i11 :: 1 bits) = ((subrange_vec_dec v__698 (( 10 :: int)::ii) (( 10 :: int)::ii) :: 1 Word.word)) in (let (i10 :: 1 bits) = ((subrange_vec_dec v__698 (( 9 :: int)::ii) (( 9 :: int)::ii) :: 1 Word.word)) in return ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec i11 ((concat_vec i4 ((concat_vec i98 ((concat_vec i10 ((concat_vec i6 ((concat_vec i7 ((concat_vec i31 ((concat_vec i5 ( 0b01 :: 2 Word.word) :: 3 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 8 Word.word)) :: 9 Word.word)) :: 11 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDIW ((v__699, rsd)) => if ((((((rsd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))) then (let (imm5 :: 1 bits) = ((subrange_vec_dec v__699 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm5 :: 1 bits) = ((subrange_vec_dec v__699 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm40 :: 5 bits) = ((subrange_vec_dec v__699 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b001 :: 3 Word.word) ((concat_vec imm5 ((concat_vec rsd ((concat_vec imm40 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LI ((v__700, rd)) => if (((rd \ zreg))) then (let (imm5 :: 1 bits) = ((subrange_vec_dec v__700 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm5 :: 1 bits) = ((subrange_vec_dec v__700 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm40 :: 5 bits) = ((subrange_vec_dec v__700 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec imm5 ((concat_vec rd ((concat_vec imm40 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDI16SP (v__701) => if ((let (nzi9 :: 1 bits) = ((subrange_vec_dec v__701 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi9 :: 1 bits) = ((subrange_vec_dec v__701 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi87 :: 2 bits) = ((subrange_vec_dec v__701 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (nzi6 :: 1 bits) = ((subrange_vec_dec v__701 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__701 (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word)) in (let (nzi4 :: 1 bits) = ((subrange_vec_dec v__701 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in (((concat_vec nzi9 ((concat_vec nzi87 ((concat_vec nzi6 ((concat_vec nzi5 nzi4 :: 2 Word.word)) :: 3 Word.word)) :: 5 Word.word)) :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))))))) then (let (nzi9 :: 1 bits) = ((subrange_vec_dec v__701 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi9 :: 1 bits) = ((subrange_vec_dec v__701 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi87 :: 2 bits) = ((subrange_vec_dec v__701 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (nzi6 :: 1 bits) = ((subrange_vec_dec v__701 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__701 (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word)) in (let (nzi4 :: 1 bits) = ((subrange_vec_dec v__701 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in return ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec nzi9 ((concat_vec ( 0b00010 :: 5 Word.word) ((concat_vec nzi4 ((concat_vec nzi6 ((concat_vec nzi87 ((concat_vec nzi5 ( 0b01 :: 2 Word.word) :: 3 Word.word)) :: 5 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LUI ((v__702, rd)) => if ((let (imm17 :: 1 bits) = ((subrange_vec_dec v__702 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__702 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__702 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in ((((rd \ zreg))) \ ((((((rd \ sp))) \ (((((concat_vec imm17 imm1612 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))))))))) then (let (imm17 :: 1 bits) = ((subrange_vec_dec v__702 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__702 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__702 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec imm17 ((concat_vec rd ((concat_vec imm1612 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SRLI ((v__703, rsd)) => if ((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__703 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__703 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__703 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))) then (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__703 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__703 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__703 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec nzui5 ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec rsd ((concat_vec nzui40 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SRAI ((v__704, rsd)) => if ((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__704 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__704 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__704 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))) then (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__704 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__704 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__704 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec nzui5 ((concat_vec ( 0b01 :: 2 Word.word) ((concat_vec rsd ((concat_vec nzui40 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ANDI ((v__705, rsd)) => (let (i5 :: 1 bits) = ((subrange_vec_dec v__705 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__705 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (i40 :: 5 bits) = ((subrange_vec_dec v__705 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec i5 ((concat_vec ( 0b10 :: 2 Word.word) ((concat_vec rsd ((concat_vec i40 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) | C_SUB ((rsd, rs2)) => return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec rsd ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec rs2 ( 0b01 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) | C_XOR ((rsd, rs2)) => return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec rsd ((concat_vec ( 0b01 :: 2 Word.word) ((concat_vec rs2 ( 0b01 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) | C_OR ((rsd, rs2)) => return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec rsd ((concat_vec ( 0b10 :: 2 Word.word) ((concat_vec rs2 ( 0b01 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) | C_AND ((rsd, rs2)) => return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec rsd ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec rs2 ( 0b01 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) | C_SUBW ((rsd, rs2)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b1 :: 1 Word.word) ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec rsd ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec rs2 ( 0b01 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDW ((rsd, rs2)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b1 :: 1 Word.word) ((concat_vec ( 0b11 :: 2 Word.word) ((concat_vec rsd ((concat_vec ( 0b01 :: 2 Word.word) ((concat_vec rs2 ( 0b01 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_J (v__706) => (let (i11 :: 1 bits) = ((subrange_vec_dec v__706 (( 10 :: int)::ii) (( 10 :: int)::ii) :: 1 Word.word)) in (let (i98 :: 2 bits) = ((subrange_vec_dec v__706 (( 8 :: int)::ii) (( 7 :: int)::ii) :: 2 Word.word)) in (let (i7 :: 1 bits) = ((subrange_vec_dec v__706 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let (i6 :: 1 bits) = ((subrange_vec_dec v__706 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__706 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (i4 :: 1 bits) = ((subrange_vec_dec v__706 (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word)) in (let (i31 :: 3 bits) = ((subrange_vec_dec v__706 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in (let (i11 :: 1 bits) = ((subrange_vec_dec v__706 (( 10 :: int)::ii) (( 10 :: int)::ii) :: 1 Word.word)) in (let (i10 :: 1 bits) = ((subrange_vec_dec v__706 (( 9 :: int)::ii) (( 9 :: int)::ii) :: 1 Word.word)) in return ((concat_vec ( 0b101 :: 3 Word.word) ((concat_vec i11 ((concat_vec i4 ((concat_vec i98 ((concat_vec i10 ((concat_vec i6 ((concat_vec i7 ((concat_vec i31 ((concat_vec i5 ( 0b01 :: 2 Word.word) :: 3 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 8 Word.word)) :: 9 Word.word)) :: 11 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))))))))) | C_BEQZ ((v__707, rs)) => (let (i8 :: 1 bits) = ((subrange_vec_dec v__707 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (i8 :: 1 bits) = ((subrange_vec_dec v__707 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (i76 :: 2 bits) = ((subrange_vec_dec v__707 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__707 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (i43 :: 2 bits) = ((subrange_vec_dec v__707 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in (let (i21 :: 2 bits) = ((subrange_vec_dec v__707 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) in return ((concat_vec ( 0b110 :: 3 Word.word) ((concat_vec i8 ((concat_vec i43 ((concat_vec rs ((concat_vec i76 ((concat_vec i21 ((concat_vec i5 ( 0b01 :: 2 Word.word) :: 3 Word.word)) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))))) | C_BNEZ ((v__708, rs)) => (let (i8 :: 1 bits) = ((subrange_vec_dec v__708 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (i8 :: 1 bits) = ((subrange_vec_dec v__708 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (i76 :: 2 bits) = ((subrange_vec_dec v__708 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__708 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (i43 :: 2 bits) = ((subrange_vec_dec v__708 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in (let (i21 :: 2 bits) = ((subrange_vec_dec v__708 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) in return ((concat_vec ( 0b111 :: 3 Word.word) ((concat_vec i8 ((concat_vec i43 ((concat_vec rs ((concat_vec i76 ((concat_vec i21 ((concat_vec i5 ( 0b01 :: 2 Word.word) :: 3 Word.word)) :: 5 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))))) | C_SLLI ((v__709, rsd)) => if ((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__709 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__709 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__709 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in ((((((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))) \ ((((((rsd \ zreg))) \ (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((nzui5 = ( 0b0 :: 1 Word.word))))))))))))))) then (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__709 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__709 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__709 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec nzui5 ((concat_vec rsd ((concat_vec nzui40 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LWSP ((v__710, rd)) => if (((rd \ zreg))) then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__710 (( 5 :: int)::ii) (( 4 :: int)::ii) :: 2 Word.word)) in (let (ui76 :: 2 bits) = ((subrange_vec_dec v__710 (( 5 :: int)::ii) (( 4 :: int)::ii) :: 2 Word.word)) in (let (ui5 :: 1 bits) = ((subrange_vec_dec v__710 (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word)) in (let (ui42 :: 3 bits) = ((subrange_vec_dec v__710 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in return ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec ui5 ((concat_vec rd ((concat_vec ui42 ((concat_vec ui76 ( 0b10 :: 2 Word.word) :: 4 Word.word)) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LDSP ((v__711, rd)) => if ((((((rd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))) then (let (ui86 :: 3 bits) = ((subrange_vec_dec v__711 (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) in (let (ui86 :: 3 bits) = ((subrange_vec_dec v__711 (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) in (let (ui5 :: 1 bits) = ((subrange_vec_dec v__711 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (ui43 :: 2 bits) = ((subrange_vec_dec v__711 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) in return ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec ui5 ((concat_vec rd ((concat_vec ui43 ((concat_vec ui86 ( 0b10 :: 2 Word.word) :: 5 Word.word)) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SWSP ((v__712, rs2)) => (let (ui76 :: 2 bits) = ((subrange_vec_dec v__712 (( 5 :: int)::ii) (( 4 :: int)::ii) :: 2 Word.word)) in (let (ui76 :: 2 bits) = ((subrange_vec_dec v__712 (( 5 :: int)::ii) (( 4 :: int)::ii) :: 2 Word.word)) in (let (ui52 :: 4 bits) = ((subrange_vec_dec v__712 (( 3 :: int)::ii) (( 0 :: int)::ii) :: 4 Word.word)) in return ((concat_vec ( 0b110 :: 3 Word.word) ((concat_vec ui52 ((concat_vec ui76 ((concat_vec rs2 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 9 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) | C_SDSP ((v__713, rs2)) => if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then (let (ui86 :: 3 bits) = ((subrange_vec_dec v__713 (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) in (let (ui86 :: 3 bits) = ((subrange_vec_dec v__713 (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__713 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in return ((concat_vec ( 0b111 :: 3 Word.word) ((concat_vec ui53 ((concat_vec ui86 ((concat_vec rs2 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 10 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_JR (rs1) => if (((rs1 \ zreg))) then return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec rs1 ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_JALR (rs1) => if (((rs1 \ zreg))) then return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b1 :: 1 Word.word) ((concat_vec rs1 ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_MV ((rd, rs2)) => if ((((((rd \ zreg))) \ (((rs2 \ zreg)))))) then return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec rd ((concat_vec rs2 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_EBREAK (_) => return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b1 :: 1 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) | C_ADD ((rsd, rs2)) => if ((((((rsd \ zreg))) \ (((rs2 \ zreg)))))) then return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b1 :: 1 Word.word) ((concat_vec rsd ((concat_vec rs2 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_NOP_HINT (v__714) => if ((let (im5 :: 1 bits) = ((subrange_vec_dec v__714 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (im5 :: 1 bits) = ((subrange_vec_dec v__714 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (im40 :: 5 bits) = ((subrange_vec_dec v__714 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((concat_vec im5 im40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))) then (let (im5 :: 1 bits) = ((subrange_vec_dec v__714 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (im5 :: 1 bits) = ((subrange_vec_dec v__714 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (im40 :: 5 bits) = ((subrange_vec_dec v__714 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec im5 ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec im40 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDI_HINT (rsd) => if (((rsd \ zreg))) then return ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec rsd ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LI_HINT (v__715) => (let (imm5 :: 1 bits) = ((subrange_vec_dec v__715 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm5 :: 1 bits) = ((subrange_vec_dec v__715 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm40 :: 5 bits) = ((subrange_vec_dec v__715 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b010 :: 3 Word.word) ((concat_vec imm5 ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec imm40 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) | C_LUI_HINT (v__716) => if ((let (imm17 :: 1 bits) = ((subrange_vec_dec v__716 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__716 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__716 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((concat_vec imm17 imm1612 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))) then (let (imm17 :: 1 bits) = ((subrange_vec_dec v__716 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__716 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__716 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec imm17 ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec imm1612 ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_MV_HINT (rs2) => if (((rs2 \ zreg))) then return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs2 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADD_HINT (rs2) => if (((rs2 \ zreg))) then return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b1 :: 1 Word.word) ((concat_vec ( 0b00000 :: 5 Word.word) ((concat_vec rs2 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SLLI_HINT ((v__717, rsd)) => if ((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__717 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__717 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__717 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((((((((concat_vec nzui5 nzui40 :: 6 Word.word)) = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))))) \ (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((nzui5 = ( 0b0 :: 1 Word.word)))))))))))) then (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__717 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__717 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__717 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in return ((concat_vec ( 0b000 :: 3 Word.word) ((concat_vec nzui5 ((concat_vec rsd ((concat_vec nzui40 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SRLI_HINT (rsd) => return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec rsd ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) | C_SRAI_HINT (rsd) => return ((concat_vec ( 0b100 :: 3 Word.word) ((concat_vec ( 0b0 :: 1 Word.word) ((concat_vec ( 0b01 :: 2 Word.word) ((concat_vec rsd ((concat_vec ( 0b00000 :: 5 Word.word) ( 0b01 :: 2 Word.word) :: 7 Word.word)) :: 10 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)) | C_FLWSP ((v__718, rd)) => and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () ))) \ ((\ (w__57 :: bool) . if w__57 then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__718 (( 5 :: int)::ii) (( 4 :: int)::ii) :: 2 Word.word)) in (let (ui76 :: 2 bits) = ((subrange_vec_dec v__718 (( 5 :: int)::ii) (( 4 :: int)::ii) :: 2 Word.word)) in (let (ui5 :: 1 bits) = ((subrange_vec_dec v__718 (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word)) in (let (ui42 :: 3 bits) = ((subrange_vec_dec v__718 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in return ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec ui5 ((concat_vec rd ((concat_vec ui42 ((concat_vec ui76 ( 0b10 :: 2 Word.word) :: 4 Word.word)) :: 7 Word.word)) :: 12 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | C_FSWSP ((v__719, rs2)) => and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () ))) \ ((\ (w__63 :: bool) . if w__63 then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__719 (( 5 :: int)::ii) (( 4 :: int)::ii) :: 2 Word.word)) in (let (ui76 :: 2 bits) = ((subrange_vec_dec v__719 (( 5 :: int)::ii) (( 4 :: int)::ii) :: 2 Word.word)) in (let (ui52 :: 4 bits) = ((subrange_vec_dec v__719 (( 3 :: int)::ii) (( 0 :: int)::ii) :: 4 Word.word)) in return ((concat_vec ( 0b111 :: 3 Word.word) ((concat_vec ui52 ((concat_vec ui76 ((concat_vec rs2 ( 0b10 :: 2 Word.word) :: 7 Word.word)) :: 9 Word.word)) :: 13 Word.word)) :: 16 Word.word))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | C_FLW ((v__720, rs1, rd)) => and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () ))) \ ((\ (w__69 :: bool) . if w__69 then (let (ui6 :: 1 bits) = ((subrange_vec_dec v__720 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (ui6 :: 1 bits) = ((subrange_vec_dec v__720 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__720 (( 3 :: int)::ii) (( 1 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__720 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in return ((concat_vec ( 0b011 :: 3 Word.word) ((concat_vec ui53 ((concat_vec rs1 ((concat_vec ui2 ((concat_vec ui6 ((concat_vec rd ( 0b00 :: 2 Word.word) :: 5 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | C_FSW ((v__721, rs1, rs2)) => and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () ))) \ ((\ (w__75 :: bool) . if w__75 then (let (ui6 :: 1 bits) = ((subrange_vec_dec v__721 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (ui6 :: 1 bits) = ((subrange_vec_dec v__721 (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__721 (( 3 :: int)::ii) (( 1 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__721 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in return ((concat_vec ( 0b111 :: 3 Word.word) ((concat_vec ui53 ((concat_vec rs1 ((concat_vec ui2 ((concat_vec ui6 ((concat_vec rs2 ( 0b00 :: 2 Word.word) :: 5 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 10 Word.word)) :: 13 Word.word)) :: 16 Word.word)))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )) | C_ILLEGAL (s) => return s | _ => assert_exp False (''Pattern match failure at unknown location'') \ exit0 () )))\ for ast :: " ast " definition encdec_compressed_backwards :: \(16)Word.word \((register_value),(ast),(exception))monad \ where \ encdec_compressed_backwards arg1 = ( (let v__722 = arg1 in if (((v__722 = ( 0x0001 :: 16 Word.word)))) then return (C_NOP () ) else if (((((let (nz96 :: 4 bits) = ((subrange_vec_dec v__722 (( 10 :: int)::ii) (( 7 :: int)::ii) :: 4 Word.word)) in (let (nz54 :: 2 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 11 :: int)::ii) :: 2 Word.word)) in (let (nz3 :: 1 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nz2 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (((concat_vec nz96 ((concat_vec nz54 ((concat_vec nz3 nz2 :: 2 Word.word)) :: 4 Word.word)) :: 8 Word.word)) \ ( 0x00 :: 8 Word.word))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))))) then (let (rd :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let (nz96 :: 4 bits) = ((subrange_vec_dec v__722 (( 10 :: int)::ii) (( 7 :: int)::ii) :: 4 Word.word)) in (let (nz54 :: 2 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 11 :: int)::ii) :: 2 Word.word)) in (let (nz3 :: 1 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nz2 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in return (C_ADDI4SPN (rd, (concat_vec nz96 ((concat_vec nz54 ((concat_vec nz3 nz2 :: 2 Word.word)) :: 4 Word.word)) :: 8 Word.word)))))))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word))))))) then (let (ui6 :: 1 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let (rs1 :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rd :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_LW ((concat_vec ui6 ((concat_vec ui53 ui2 :: 4 Word.word)) :: 5 Word.word), rs1, rd))))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))))) then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (rs1 :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rd :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_LD ((concat_vec ui76 ui53 :: 5 Word.word), rs1, rd)))))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b110 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word))))))) then (let (ui6 :: 1 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let (rs2 :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let (rs1 :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in return (C_SW ((concat_vec ui6 ((concat_vec ui53 ui2 :: 4 Word.word)) :: 5 Word.word), rs1, rs2))))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))))) then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: 3 bits) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let (rs1 :: 3 bits) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in return (C_SD ((concat_vec ui76 ui53 :: 5 Word.word), rs1, rs2)))))) else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzi40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in ((((((concat_vec nzi5 nzi40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzi40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_ADDI ((concat_vec nzi5 nzi40 :: 6 Word.word), rsd))))) else if (((((((( 32 :: int)::ii) = (( 32 :: int)::ii)))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (i98 :: 2 bits) = ((subrange_vec_dec v__722 (( 10 :: int)::ii) (( 9 :: int)::ii) :: 2 Word.word)) in (let (i7 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let (i6 :: 1 bits) = ((subrange_vec_dec v__722 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__722 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (i4 :: 1 bits) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in (let (i31 :: 3 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) in (let (i11 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (i10 :: 1 bits) = ((subrange_vec_dec v__722 (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word)) in return (C_JAL ((concat_vec i11 ((concat_vec i10 ((concat_vec i98 ((concat_vec i7 ((concat_vec i6 ((concat_vec i5 ((concat_vec i4 i31 :: 4 Word.word)) :: 5 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 9 Word.word)) :: 10 Word.word)) :: 11 Word.word))))))))))) else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in ((((rsd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_ADDIW ((concat_vec imm5 imm40 :: 6 Word.word), rsd))))) else if (((((let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rd \ zreg))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_LI ((concat_vec imm5 imm40 :: 6 Word.word), rd))))) else if (((((let (nzi9 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzi87 :: 2 bits) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (nzi6 :: 1 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__722 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (nzi4 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (((concat_vec nzi9 ((concat_vec nzi87 ((concat_vec nzi6 ((concat_vec nzi5 nzi4 :: 2 Word.word)) :: 3 Word.word)) :: 5 Word.word)) :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) = ( 0b00010 :: 5 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then (let (nzi9 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzi87 :: 2 bits) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (nzi6 :: 1 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__722 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (nzi4 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in return (C_ADDI16SP ((concat_vec nzi9 ((concat_vec nzi87 ((concat_vec nzi6 ((concat_vec nzi5 nzi4 :: 2 Word.word)) :: 3 Word.word)) :: 5 Word.word)) :: 6 Word.word)))))))) else if (((((let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in ((((rd \ zreg))) \ ((((((rd \ sp))) \ (((((concat_vec imm17 imm1612 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_LUI ((concat_vec imm17 imm1612 :: 6 Word.word), rd))))) else if (((((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b100 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_SRLI ((concat_vec nzui5 nzui40 :: 6 Word.word), rsd))))) else if (((((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b100 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_SRAI ((concat_vec nzui5 nzui40 :: 6 Word.word), rsd))))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b100 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (i40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_ANDI ((concat_vec i5 i40 :: 6 Word.word), rsd))))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100011 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_SUB (rsd, rs2)))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100011 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_XOR (rsd, rs2)))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100011 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_OR (rsd, rs2)))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100011 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_AND (rsd, rs2)))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100111 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_SUBW (rsd, rs2)))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100111 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_ADDW (rsd, rs2)))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then (let (i98 :: 2 bits) = ((subrange_vec_dec v__722 (( 10 :: int)::ii) (( 9 :: int)::ii) :: 2 Word.word)) in (let (i7 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let (i6 :: 1 bits) = ((subrange_vec_dec v__722 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__722 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (i4 :: 1 bits) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in (let (i31 :: 3 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) in (let (i11 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (i10 :: 1 bits) = ((subrange_vec_dec v__722 (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word)) in return (C_J ((concat_vec i11 ((concat_vec i10 ((concat_vec i98 ((concat_vec i7 ((concat_vec i6 ((concat_vec i5 ((concat_vec i4 i31 :: 4 Word.word)) :: 5 Word.word)) :: 6 Word.word)) :: 7 Word.word)) :: 9 Word.word)) :: 10 Word.word)) :: 11 Word.word))))))))))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b110 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then (let (rs :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (i8 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (i76 :: 2 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__722 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (i43 :: 2 bits) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) in (let (i21 :: 2 bits) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in return (C_BEQZ ((concat_vec i8 ((concat_vec i76 ((concat_vec i5 ((concat_vec i43 i21 :: 4 Word.word)) :: 5 Word.word)) :: 7 Word.word)) :: 8 Word.word), rs)))))))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then (let (rs :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (i8 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (i76 :: 2 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__722 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (i43 :: 2 bits) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) in (let (i21 :: 2 bits) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in return (C_BNEZ ((concat_vec i8 ((concat_vec i76 ((concat_vec i5 ((concat_vec i43 i21 :: 4 Word.word)) :: 5 Word.word)) :: 7 Word.word)) :: 8 Word.word), rs)))))))) else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in ((((((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))) \ ((((((rsd \ zreg))) \ (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((nzui5 = ( 0b0 :: 1 Word.word))))))))))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_SLLI ((concat_vec nzui5 nzui40 :: 6 Word.word), rsd))))) else if (((((let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rd \ zreg))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__722 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in (let (ui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (ui42 :: 3 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 4 :: int)::ii) :: 3 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (C_LWSP ((concat_vec ui76 ((concat_vec ui5 ui42 :: 4 Word.word)) :: 6 Word.word), rd)))))) else if (((((let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in ((((rd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (ui86 :: 3 bits) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let (ui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (ui43 :: 2 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (C_LDSP ((concat_vec ui86 ((concat_vec ui5 ui43 :: 3 Word.word)) :: 6 Word.word), rd)))))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b110 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word))))))) then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__722 (( 8 :: int)::ii) (( 7 :: int)::ii) :: 2 Word.word)) in (let (ui52 :: 4 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 9 :: int)::ii) :: 4 Word.word)) in (let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_SWSP ((concat_vec ui76 ui52 :: 6 Word.word), rs2))))) else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (ui86 :: 3 bits) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_SDSP ((concat_vec ui86 ui53 :: 6 Word.word), rs2))))) else if (((((let (rs1 :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rs1 \ zreg))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x8 :: 4 Word.word)))) \ (((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000010 :: 7 Word.word)))))))))) then (let (rs1 :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (C_JR rs1)) else if (((((let (rs1 :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rs1 \ zreg))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x9 :: 4 Word.word)))) \ (((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000010 :: 7 Word.word)))))))))) then (let (rs1 :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (C_JALR rs1)) else if (((((let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in ((((rd \ zreg))) \ (((rs2 \ zreg))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x8 :: 4 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (C_MV (rd, rs2)))) else if (((v__722 = ( 0x9002 :: 16 Word.word)))) then return (C_EBREAK () ) else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in ((((rsd \ zreg))) \ (((rs2 \ zreg))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x9 :: 4 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_ADD (rsd, rs2)))) else if (((((let (im5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (im40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((concat_vec im5 im40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) = ( 0b00000 :: 5 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then (let (im5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (im40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_NOP_HINT ((concat_vec im5 im40 :: 6 Word.word))))) else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rsd \ zreg))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x0 :: 4 Word.word)))) \ (((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))))))))) then (let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (C_ADDI_HINT rsd)) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) = ( 0b00000 :: 5 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then (let (imm5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_LI_HINT ((concat_vec imm5 imm40 :: 6 Word.word))))) else if (((((let (imm17 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((concat_vec imm17 imm1612 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) = ( 0b00000 :: 5 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then (let (imm17 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_LUI_HINT ((concat_vec imm17 imm1612 :: 6 Word.word))))) else if (((((let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (rs2 \ zreg))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 7 :: int)::ii) :: 9 Word.word)) = ( 0b100000000 :: 9 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_MV_HINT rs2)) else if (((((let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (rs2 \ zreg))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 7 :: int)::ii) :: 9 Word.word)) = ( 0b100100000 :: 9 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_ADD_HINT rs2)) else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((((((((concat_vec nzui5 nzui40 :: 6 Word.word)) = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))))) \ (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((nzui5 = ( 0b0 :: 1 Word.word)))))))))))) \ ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then (let (rsd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_SLLI_HINT ((concat_vec nzui5 nzui40 :: 6 Word.word), rsd))))) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100000 :: 6 Word.word)))) \ (((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in return (C_SRLI_HINT rsd)) else if ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100001 :: 6 Word.word)))) \ (((((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word))))))) then (let (rsd :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in return (C_SRAI_HINT rsd)) else and_boolM (and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () )))) (return ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))) \ ((\ (w__4 :: bool) . if w__4 then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__722 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in (let (ui5 :: 1 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (ui42 :: 3 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 4 :: int)::ii) :: 3 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__722 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in return (C_FLWSP ((concat_vec ui76 ((concat_vec ui5 ui42 :: 4 Word.word)) :: 6 Word.word), rd)))))) else and_boolM (and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () )))) (return ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))) \ ((\ (w__9 :: bool) . if w__9 then (let (ui76 :: 2 bits) = ((subrange_vec_dec v__722 (( 8 :: int)::ii) (( 7 :: int)::ii) :: 2 Word.word)) in (let (ui52 :: 4 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 9 :: int)::ii) :: 4 Word.word)) in (let (rs2 :: regidx) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in return (C_FSWSP ((concat_vec ui76 ui52 :: 6 Word.word), rs2))))) else and_boolM (and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () )))) (return ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))) \ ((\ (w__14 :: bool) . if w__14 then (let (ui6 :: 1 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let (rs1 :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rd :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in return (C_FLW ((concat_vec ui6 ((concat_vec ui53 ui2 :: 4 Word.word)) :: 5 Word.word), rs1, rd))))))) else and_boolM (and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () )))) (return ((((((((subrange_vec_dec v__722 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__722 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))) \ ((\ (w__19 :: bool) . return (if w__19 then (let (ui6 :: 1 bits) = ((subrange_vec_dec v__722 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__722 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__722 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let (rs2 :: cregidx) = ((subrange_vec_dec v__722 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let (rs1 :: cregidx) = ((subrange_vec_dec v__722 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in C_FSW ((concat_vec ui6 ((concat_vec ui53 ui2 :: 4 Word.word)) :: 5 Word.word), rs1, rs2)))))) else C_ILLEGAL v__722)))))))))))\ for arg1 :: "(16)Word.word " definition encdec_compressed_forwards_matches :: \ ast \((register_value),(bool),(exception))monad \ where \ encdec_compressed_forwards_matches ast = ( (let arg1 = ast in (case arg1 of C_NOP (_) => return True | C_ADDI4SPN ((rd, v__912)) => return (if ((let (nz96 :: 4 bits) = ((subrange_vec_dec v__912 (( 7 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word)) in (let (nz96 :: 4 bits) = ((subrange_vec_dec v__912 (( 7 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word)) in (let (nz54 :: 2 bits) = ((subrange_vec_dec v__912 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in (let (nz3 :: 1 bits) = ((subrange_vec_dec v__912 (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word)) in (let (nz2 :: 1 bits) = ((subrange_vec_dec v__912 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in (((concat_vec nz96 ((concat_vec nz54 ((concat_vec nz3 nz2 :: 2 Word.word)) :: 4 Word.word)) :: 8 Word.word)) \ ( 0x00 :: 8 Word.word)))))))) then True else False) | C_LW ((v__913, rs1, rd)) => return True | C_LD ((v__914, rs1, rd)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | C_SW ((v__915, rs1, rs2)) => return True | C_SD ((v__916, rs1, rs2)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | C_ADDI ((v__917, rsd)) => return (if ((let (nzi5 :: 1 bits) = ((subrange_vec_dec v__917 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__917 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi40 :: 5 bits) = ((subrange_vec_dec v__917 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in ((((((concat_vec nzi5 nzi40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))))))) then True else False) | C_JAL (v__918) => return (if ((((( 32 :: int)::ii) = (( 32 :: int)::ii)))) then True else False) | C_ADDIW ((v__919, rsd)) => return (if ((((((rsd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))) then True else False) | C_LI ((v__920, rd)) => return (if (((rd \ zreg))) then True else False) | C_ADDI16SP (v__921) => return (if ((let (nzi9 :: 1 bits) = ((subrange_vec_dec v__921 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi9 :: 1 bits) = ((subrange_vec_dec v__921 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi87 :: 2 bits) = ((subrange_vec_dec v__921 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (nzi6 :: 1 bits) = ((subrange_vec_dec v__921 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__921 (( 1 :: int)::ii) (( 1 :: int)::ii) :: 1 Word.word)) in (let (nzi4 :: 1 bits) = ((subrange_vec_dec v__921 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in (((concat_vec nzi9 ((concat_vec nzi87 ((concat_vec nzi6 ((concat_vec nzi5 nzi4 :: 2 Word.word)) :: 3 Word.word)) :: 5 Word.word)) :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))))))) then True else False) | C_LUI ((v__922, rd)) => return (if ((let (imm17 :: 1 bits) = ((subrange_vec_dec v__922 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__922 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__922 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in ((((rd \ zreg))) \ ((((((rd \ sp))) \ (((((concat_vec imm17 imm1612 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))))))))) then True else False) | C_SRLI ((v__923, rsd)) => return (if ((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__923 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__923 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__923 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))) then True else False) | C_SRAI ((v__924, rsd)) => return (if ((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__924 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__924 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__924 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))) then True else False) | C_ANDI ((v__925, rsd)) => return True | C_SUB ((rsd, rs2)) => return True | C_XOR ((rsd, rs2)) => return True | C_OR ((rsd, rs2)) => return True | C_AND ((rsd, rs2)) => return True | C_SUBW ((rsd, rs2)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | C_ADDW ((rsd, rs2)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | C_J (v__926) => return True | C_BEQZ ((v__927, rs)) => return True | C_BNEZ ((v__928, rs)) => return True | C_SLLI ((v__929, rsd)) => return (if ((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__929 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__929 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__929 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in ((((((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))) \ ((((((rsd \ zreg))) \ (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((nzui5 = ( 0b0 :: 1 Word.word))))))))))))))) then True else False) | C_LWSP ((v__930, rd)) => return (if (((rd \ zreg))) then True else False) | C_LDSP ((v__931, rd)) => return (if ((((((rd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))) then True else False) | C_SWSP ((v__932, rs2)) => return True | C_SDSP ((v__933, rs2)) => return (if ((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) then True else False) | C_JR (rs1) => return (if (((rs1 \ zreg))) then True else False) | C_JALR (rs1) => return (if (((rs1 \ zreg))) then True else False) | C_MV ((rd, rs2)) => return (if ((((((rd \ zreg))) \ (((rs2 \ zreg)))))) then True else False) | C_EBREAK (_) => return True | C_ADD ((rsd, rs2)) => return (if ((((((rsd \ zreg))) \ (((rs2 \ zreg)))))) then True else False) | C_NOP_HINT (v__934) => return (if ((let (im5 :: 1 bits) = ((subrange_vec_dec v__934 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (im5 :: 1 bits) = ((subrange_vec_dec v__934 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (im40 :: 5 bits) = ((subrange_vec_dec v__934 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((concat_vec im5 im40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))) then True else False) | C_ADDI_HINT (rsd) => return (if (((rsd \ zreg))) then True else False) | C_LI_HINT (v__935) => return True | C_LUI_HINT (v__936) => return (if ((let (imm17 :: 1 bits) = ((subrange_vec_dec v__936 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__936 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__936 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((concat_vec imm17 imm1612 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))) then True else False) | C_MV_HINT (rs2) => return (if (((rs2 \ zreg))) then True else False) | C_ADD_HINT (rs2) => return (if (((rs2 \ zreg))) then True else False) | C_SLLI_HINT ((v__937, rsd)) => return (if ((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__937 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__937 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__937 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (((((((((concat_vec nzui5 nzui40 :: 6 Word.word)) = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))))) \ (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((nzui5 = ( 0b0 :: 1 Word.word)))))))))))) then True else False) | C_SRLI_HINT (rsd) => return True | C_SRAI_HINT (rsd) => return True | C_FLWSP ((v__938, rd)) => and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () ))) \ ((\ (w__3 :: bool) . return (if w__3 then True else False))) | C_FSWSP ((v__939, rs2)) => and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () ))) \ ((\ (w__7 :: bool) . return (if w__7 then True else False))) | C_FLW ((v__940, rs1, rd)) => and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () ))) \ ((\ (w__11 :: bool) . return (if w__11 then True else False))) | C_FSW ((v__941, rs1, rs2)) => and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () ))) \ ((\ (w__15 :: bool) . return (if w__15 then True else False))) | C_ILLEGAL (s) => return True | _ => return False )))\ for ast :: " ast " definition encdec_compressed_backwards_matches :: \(16)Word.word \((register_value),(bool),(exception))monad \ where \ encdec_compressed_backwards_matches arg1 = ( (let v__942 = arg1 in if (((v__942 = ( 0x0001 :: 16 Word.word)))) then return True else if (((((let (nz96 :: 4 bits) = ((subrange_vec_dec v__942 (( 10 :: int)::ii) (( 7 :: int)::ii) :: 4 Word.word)) in (let (nz54 :: 2 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 11 :: int)::ii) :: 2 Word.word)) in (let (nz3 :: 1 bits) = ((subrange_vec_dec v__942 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nz2 :: 1 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (((concat_vec nz96 ((concat_vec nz54 ((concat_vec nz3 nz2 :: 2 Word.word)) :: 4 Word.word)) :: 8 Word.word)) \ ( 0x00 :: 8 Word.word))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b110 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))))) then return True else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzi40 :: 5 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in ((((((concat_vec nzi5 nzi40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 32 :: int)::ii)))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in ((((rsd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b001 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if (((((let (rd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rd \ zreg))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if (((((let (nzi9 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzi87 :: 2 bits) = ((subrange_vec_dec v__942 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (nzi6 :: 1 bits) = ((subrange_vec_dec v__942 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__942 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in (let (nzi4 :: 1 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (((concat_vec nzi9 ((concat_vec nzi87 ((concat_vec nzi6 ((concat_vec nzi5 nzi4 :: 2 Word.word)) :: 3 Word.word)) :: 5 Word.word)) :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) = ( 0b00010 :: 5 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then return True else if (((((let (rd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (imm17 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in ((((rd \ zreg))) \ ((((((rd \ sp))) \ (((((concat_vec imm17 imm1612 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))))))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if (((((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b100 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then return True else if (((((let (nzui5 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b100 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b100 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100011 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100011 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100011 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100011 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100111 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100111 :: 6 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b101 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b110 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then return True else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in ((((((concat_vec nzui5 nzui40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word)))) \ ((((((rsd \ zreg))) \ (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((nzui5 = ( 0b0 :: 1 Word.word))))))))))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if (((((let (rd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rd \ zreg))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if (((((let (rd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in ((((rd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b110 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word))))))) then return True else if (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if (((((let (rs1 :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rs1 \ zreg))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x8 :: 4 Word.word)))) \ (((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000010 :: 7 Word.word)))))))))) then return True else if (((((let (rs1 :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rs1 \ zreg))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x9 :: 4 Word.word)))) \ (((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000010 :: 7 Word.word)))))))))) then return True else if (((((let (rs2 :: regidx) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let (rd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in ((((rd \ zreg))) \ (((rs2 \ zreg))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x8 :: 4 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if (((v__942 = ( 0x9002 :: 16 Word.word)))) then return True else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: regidx) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in ((((rsd \ zreg))) \ (((rs2 \ zreg))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x9 :: 4 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if (((((let (im5 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (im40 :: 5 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((concat_vec im5 im40 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) = ( 0b00000 :: 5 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then return True else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (rsd \ zreg))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = ( 0x0 :: 4 Word.word)))) \ (((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b010 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) = ( 0b00000 :: 5 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))))))))) then return True else if (((((let (imm17 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((concat_vec imm17 imm1612 :: 6 Word.word)) \ ( 0b000000 :: 6 Word.word))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ ((((((((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) = ( 0b00000 :: 5 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))))))))) then return True else if (((((let (rs2 :: regidx) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (rs2 \ zreg))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 7 :: int)::ii) :: 9 Word.word)) = ( 0b100000000 :: 9 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if (((((let (rs2 :: regidx) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (rs2 \ zreg))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 7 :: int)::ii) :: 9 Word.word)) = ( 0b100100000 :: 9 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if (((((let (rsd :: regidx) = ((subrange_vec_dec v__942 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__942 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (((((((((concat_vec nzui5 nzui40 :: 6 Word.word)) = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))))) \ (((((((( 32 :: int)::ii) = (( 64 :: int)::ii)))) \ (((nzui5 = ( 0b0 :: 1 Word.word)))))))))))) \ ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100000 :: 6 Word.word)))) \ (((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word))))))) then return True else if ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = ( 0b100001 :: 6 Word.word)))) \ (((((subrange_vec_dec v__942 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = ( 0b0000001 :: 7 Word.word))))))) then return True else and_boolM (and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () )))) (return ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))) \ ((\ (w__4 :: bool) . if w__4 then return True else and_boolM (and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () )))) (return ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))))))) \ ((\ (w__9 :: bool) . if w__9 then return True else and_boolM (and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () )))) (return ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b011 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))) \ ((\ (w__14 :: bool) . if w__14 then return True else and_boolM (and_boolM (return ((((( 32 :: int)::ii) = (( 32 :: int)::ii))))) (and_boolM ((haveRVC () )) ((haveFExt () )))) (return ((((((((subrange_vec_dec v__942 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = ( 0b111 :: 3 Word.word)))) \ (((((subrange_vec_dec v__942 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))))))) \ ((\ (w__19 :: bool) . return (if w__19 then True else True)))))))))))\ for arg1 :: "(16)Word.word " \ \\val execute_WFI : unit -> M Retired\\ definition execute_WFI :: \ unit \((register_value),(Retired),(exception))monad \ where \ execute_WFI _ = ( read_reg cur_privilege_ref \ ((\ (w__0 :: Privilege) . (case w__0 of Machine => platform_wfi () \ return RETIRE_SUCCESS | Supervisor => read_reg mstatus_ref \ ((\ (w__1 :: Mstatus) . if (((((get_Mstatus_TW w__1 :: 1 Word.word)) = ( 0b1 :: 1 Word.word)))) then handle_illegal () \ return RETIRE_FAIL else platform_wfi () \ return RETIRE_SUCCESS)) | User => handle_illegal () \ return RETIRE_FAIL ))))\ \ \\val execute_UTYPE : mword ty20 -> mword ty5 -> uop -> M Retired\\ definition execute_UTYPE :: \(20)Word.word \(5)Word.word \ uop \((register_value),(Retired),(exception))monad \ where \ execute_UTYPE imm rd op1 = ( (let (off :: xlenbits) = ((EXTS (( 32 :: int)::ii) ((concat_vec imm ( 0x000 :: 12 Word.word) :: 32 Word.word)) :: 32 Word.word)) in (case op1 of RISCV_LUI => return off | RISCV_AUIPC => (get_arch_pc () :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return ((add_vec w__0 off :: 32 Word.word)))) ) \ ((\ (ret :: xlenbits) . wX_bits rd ret \ return RETIRE_SUCCESS))))\ for imm :: "(20)Word.word " and rd :: "(5)Word.word " and op1 :: " uop " \ \\val execute_URET : unit -> M Retired\\ definition execute_URET :: \ unit \((register_value),(Retired),(exception))monad \ where \ execute_URET _ = ( or_boolM (haveUsrMode () \ ((\ (w__0 :: bool) . return ((\ w__0))))) (return ((\ ((sys_enable_next () ))))) \ ((\ (w__1 :: bool) . (if w__1 then handle_illegal () else if ((\ ((ext_check_xret_priv User)))) then return ((ext_fail_xret_priv () )) else read_reg cur_privilege_ref \ ((\ (w__2 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__3 :: 32 Word.word) . (exception_handler w__2 (CTL_URET () ) w__3 :: ( 32 Word.word) M) \ ((\ (w__4 :: 32 Word.word) . set_next_pc w__4))))))) \ return RETIRE_FAIL)))\ \ \\val execute_STORE_FP : mword ty12 -> mword ty5 -> mword ty5 -> word_width -> M Retired\\ definition execute_STORE_FP :: \(12)Word.word \(5)Word.word \(5)Word.word \ word_width \((register_value),(Retired),(exception))monad \ where \ execute_STORE_FP imm rs2 rs1 width = ( (let (offset :: xlenbits) = ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) in (let (aq, rl, con) = (False, False, False) in ext_data_get_addr rs1 offset (Write Data) width \ ((\ (w__0 :: unit Ext_DataAddr_Check) . (case w__0 of Ext_DataAddr_Error (e) => (let (_ :: unit) = (ext_handle_data_check_error e) in return RETIRE_FAIL) | Ext_DataAddr_OK (vaddr) => if ((check_misaligned vaddr width)) then handle_mem_exception vaddr (E_SAMO_Addr_Align () ) \ return RETIRE_FAIL else (translateAddr vaddr (Write Data) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__1 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__1 of TR_Failure ((e, _)) => handle_mem_exception vaddr e \ return RETIRE_FAIL | TR_Address ((addr, _)) => (case width of BYTE => return (MemValue () ) | HALF => return (MemValue () ) | WORD => mem_write_ea addr (( 4 :: int)::ii) aq rl False | DOUBLE => mem_write_ea addr (( 8 :: int)::ii) aq rl False ) \ ((\ (eares :: unit MemoryOpResult) . (case eares of MemException (e) => handle_mem_exception addr e \ return RETIRE_FAIL | MemValue (_) => (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ rs2_val . (case (width, (( 32 :: int)::ii)) of (BYTE, g__321) => handle_illegal () \ return RETIRE_FAIL | (HALF, g__322) => handle_illegal () \ return RETIRE_FAIL | (WORD, g__323) => mem_write_value addr (( 4 :: int)::ii) ((subrange_vec_dec rs2_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) aq rl con \ ((\ (w__4 :: bool MemoryOpResult) . process_fstore vaddr w__4)) | _ => assert_exp False (''Pattern match failure at model/riscv_insts_fext.sail 401:14 - 406:15'') \ exit0 () ))) ))) ))) ))))))\ for imm :: "(12)Word.word " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and width :: " word_width " \ \\val execute_STORECON : bool -> bool -> mword ty5 -> mword ty5 -> word_width -> mword ty5 -> M Retired\\ definition execute_STORECON :: \ bool \ bool \(5)Word.word \(5)Word.word \ word_width \(5)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_STORECON aq rl rs2 rs1 width rd = ( speculate_conditional_success () \ ((\ (w__0 :: bool) . if (((w__0 = False))) then wX_bits rd ((EXTZ (( 32 :: int)::ii) ( 0b1 :: 1 Word.word) :: 32 Word.word)) \ return RETIRE_SUCCESS else haveAtomics () \ ((\ (w__1 :: bool) . if w__1 then ext_data_get_addr rs1 ((zeros_implicit (( 32 :: int)::ii) :: 32 Word.word)) (Write Data) width \ ((\ (w__2 :: unit Ext_DataAddr_Check) . (case w__2 of Ext_DataAddr_Error (e) => (let (_ :: unit) = (ext_handle_data_check_error e) in return RETIRE_FAIL) | Ext_DataAddr_OK (vaddr) => (let (aligned :: bool) = ((case width of BYTE => True | HALF => (((subrange_vec_dec vaddr (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)) | WORD => (((subrange_vec_dec vaddr (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)) | DOUBLE => (((subrange_vec_dec vaddr (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)) )) in if ((\ aligned)) then handle_mem_exception vaddr (E_SAMO_Addr_Align () ) \ return RETIRE_FAIL else if (((((match_reservation vaddr)) = False))) then wX_bits rd ((EXTZ (( 32 :: int)::ii) ( 0b1 :: 1 Word.word) :: 32 Word.word)) \ ((let (_ :: unit) = (cancel_reservation () ) in return RETIRE_SUCCESS)) else (translateAddr vaddr (Write Data) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__3 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__3 of TR_Failure ((e, _)) => handle_mem_exception vaddr e \ return RETIRE_FAIL | TR_Address ((addr, _)) => (case (width, (( 32 :: int)::ii)) of (BYTE, _) => mem_write_ea addr (( 1 :: int)::ii) (((aq \ rl))) rl True | (HALF, _) => mem_write_ea addr (( 2 :: int)::ii) (((aq \ rl))) rl True | (WORD, _) => mem_write_ea addr (( 4 :: int)::ii) (((aq \ rl))) rl True | _ => internal_error (''STORECON expected word or double'') ) \ ((\ (eares :: unit MemoryOpResult) . (case eares of MemException (e) => handle_mem_exception addr e \ return RETIRE_FAIL | MemValue (_) => (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ rs2_val . (case (width, (( 32 :: int)::ii)) of (BYTE, _) => mem_write_value addr (( 1 :: int)::ii) ((subrange_vec_dec rs2_val (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)) (((aq \ rl))) rl True | (HALF, _) => mem_write_value addr (( 2 :: int)::ii) ((subrange_vec_dec rs2_val (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) (((aq \ rl))) rl True | (WORD, _) => mem_write_value addr (( 4 :: int)::ii) ((subrange_vec_dec rs2_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) (((aq \ rl))) rl True | _ => internal_error (''STORECON expected word or double'') ) \ ((\ (res :: bool MemoryOpResult) . (case res of MemValue (True) => wX_bits rd ((EXTZ (( 32 :: int)::ii) ( 0b0 :: 1 Word.word) :: 32 Word.word)) \ ((let (_ :: unit) = (cancel_reservation () ) in return RETIRE_SUCCESS)) | MemValue (False) => wX_bits rd ((EXTZ (( 32 :: int)::ii) ( 0b1 :: 1 Word.word) :: 32 Word.word)) \ ((let (_ :: unit) = (cancel_reservation () ) in return RETIRE_SUCCESS)) | MemException (e) => handle_mem_exception addr e \ return RETIRE_FAIL ))))) ))) )))) ))) else handle_illegal () \ return RETIRE_FAIL)))))\ for aq :: " bool " and rl :: " bool " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and width :: " word_width " and rd :: "(5)Word.word " \ \\val execute_STORE : mword ty12 -> mword ty5 -> mword ty5 -> word_width -> bool -> bool -> M Retired\\ definition execute_STORE :: \(12)Word.word \(5)Word.word \(5)Word.word \ word_width \ bool \ bool \((register_value),(Retired),(exception))monad \ where \ execute_STORE imm rs2 rs1 width aq rl = ( (let (offset :: xlenbits) = ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) in ext_data_get_addr rs1 offset (Write Data) width \ ((\ (w__0 :: unit Ext_DataAddr_Check) . (case w__0 of Ext_DataAddr_Error (e) => (let (_ :: unit) = (ext_handle_data_check_error e) in return RETIRE_FAIL) | Ext_DataAddr_OK (vaddr) => if ((check_misaligned vaddr width)) then handle_mem_exception vaddr (E_SAMO_Addr_Align () ) \ return RETIRE_FAIL else (translateAddr vaddr (Write Data) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__1 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__1 of TR_Failure ((e, _)) => handle_mem_exception vaddr e \ return RETIRE_FAIL | TR_Address ((paddr, _)) => (case width of BYTE => mem_write_ea paddr (( 1 :: int)::ii) aq rl False | HALF => mem_write_ea paddr (( 2 :: int)::ii) aq rl False | WORD => mem_write_ea paddr (( 4 :: int)::ii) aq rl False | DOUBLE => mem_write_ea paddr (( 8 :: int)::ii) aq rl False ) \ ((\ (eares :: unit MemoryOpResult) . (case eares of MemException (e) => handle_mem_exception vaddr e \ return RETIRE_FAIL | MemValue (_) => (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ rs2_val . (case (width, (( 32 :: int)::ii)) of (BYTE, g__313) => mem_write_value paddr (( 1 :: int)::ii) ((subrange_vec_dec rs2_val (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)) aq rl False | (HALF, g__314) => mem_write_value paddr (( 2 :: int)::ii) ((subrange_vec_dec rs2_val (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) aq rl False | (WORD, g__315) => mem_write_value paddr (( 4 :: int)::ii) ((subrange_vec_dec rs2_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) aq rl False | _ => assert_exp False (''Pattern match failure at model/riscv_insts_base.sail 394:47 - 399:15'') \ exit0 () ) \ ((\ (res :: bool MemoryOpResult) . (case res of MemValue (True) => return RETIRE_SUCCESS | MemValue (False) => internal_error (''store got false from mem_write_value'') | MemException (e) => handle_mem_exception vaddr e \ return RETIRE_FAIL ))))) ))) ))) )))))\ for imm :: "(12)Word.word " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and width :: " word_width " and aq :: " bool " and rl :: " bool " \ \\val execute_SRET : unit -> M Retired\\ definition execute_SRET :: \ unit \((register_value),(Retired),(exception))monad \ where \ execute_SRET _ = ( read_reg cur_privilege_ref \ ((\ (w__0 :: Privilege) . (case w__0 of User => return True | Supervisor => or_boolM (haveSupMode () \ ((\ (w__1 :: bool) . return ((\ w__1))))) (read_reg mstatus_ref \ ((\ (w__2 :: Mstatus) . return (((((get_Mstatus_TSR w__2 :: 1 Word.word)) = ( 0b1 :: 1 Word.word))))))) | Machine => haveSupMode () \ ((\ (w__4 :: bool) . return ((\ w__4)))) ) \ ((\ (sret_illegal :: bool) . if sret_illegal then handle_illegal () \ return RETIRE_FAIL else if ((\ ((ext_check_xret_priv Supervisor)))) then (let (_ :: unit) = (ext_fail_xret_priv () ) in return RETIRE_FAIL) else read_reg cur_privilege_ref \ ((\ (w__5 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__6 :: 32 Word.word) . (exception_handler w__5 (CTL_SRET () ) w__6 :: ( 32 Word.word) M) \ ((\ (w__7 :: 32 Word.word) . set_next_pc w__7 \ return RETIRE_SUCCESS)))))))))))\ \ \\val execute_SHIFTW : mword ty5 -> mword ty5 -> mword ty5 -> sop -> M Retired\\ definition execute_SHIFTW :: \(5)Word.word \(5)Word.word \(5)Word.word \ sop \((register_value),(Retired),(exception))monad \ where \ execute_SHIFTW shamt rs1 rd op1 = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val = ((subrange_vec_dec w__0 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (let (result :: 32 bits) = ((case op1 of RISCV_SLLI => (shift_bits_left rs1_val shamt :: 32 Word.word) | RISCV_SRLI => (shift_bits_right rs1_val shamt :: 32 Word.word) | RISCV_SRAI => (shift_right_arith32 rs1_val shamt :: 32 Word.word) )) in wX_bits rd ((EXTS (( 32 :: int)::ii) result :: 32 Word.word)) \ return RETIRE_SUCCESS)))))\ for shamt :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and op1 :: " sop " \ \\val execute_SHIFTIWOP : mword ty5 -> mword ty5 -> mword ty5 -> sopw -> M Retired\\ definition execute_SHIFTIWOP :: \(5)Word.word \(5)Word.word \(5)Word.word \ sopw \((register_value),(Retired),(exception))monad \ where \ execute_SHIFTIWOP shamt rs1 rd op1 = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val . (let (result :: 32 bits) = ((case op1 of RISCV_SLLIW => (shift_bits_left ((subrange_vec_dec rs1_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) shamt :: 32 Word.word) | RISCV_SRLIW => (shift_bits_right ((subrange_vec_dec rs1_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) shamt :: 32 Word.word) | RISCV_SRAIW => (shift_right_arith32 ((subrange_vec_dec rs1_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) shamt :: 32 Word.word) )) in wX_bits rd ((EXTS (( 32 :: int)::ii) result :: 32 Word.word)) \ return RETIRE_SUCCESS))))\ for shamt :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and op1 :: " sopw " \ \\val execute_SHIFTIOP : mword ty6 -> mword ty5 -> mword ty5 -> sop -> M Retired\\ definition execute_SHIFTIOP :: \(6)Word.word \(5)Word.word \(5)Word.word \ sop \((register_value),(Retired),(exception))monad \ where \ execute_SHIFTIOP shamt rs1 rd op1 = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val . (let (result :: xlenbits) = ((case op1 of RISCV_SLLI => (shift_bits_left rs1_val ((subrange_vec_dec shamt (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) | RISCV_SRLI => (shift_bits_right rs1_val ((subrange_vec_dec shamt (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) | RISCV_SRAI => (shift_right_arith32 rs1_val ((subrange_vec_dec shamt (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) )) in wX_bits rd result \ return RETIRE_SUCCESS))))\ for shamt :: "(6)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and op1 :: " sop " \ \\val execute_SFENCE_VMA : mword ty5 -> mword ty5 -> M Retired\\ definition execute_SFENCE_VMA :: \(5)Word.word \(5)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_SFENCE_VMA rs1 rs2 = ( (if (((rs1 = ( 0b00000 :: 5 Word.word)))) then return None else (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . return (Some w__0)))) \ ((\ (addr :: xlenbits option) . (if (((rs2 = ( 0b00000 :: 5 Word.word)))) then return None else (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . return (Some w__1)))) \ ((\ (asid :: xlenbits option) . read_reg cur_privilege_ref \ ((\ (w__2 :: Privilege) . (case w__2 of User => handle_illegal () \ return RETIRE_FAIL | Supervisor => read_reg mstatus_ref \ ((\ (w__3 :: Mstatus) . read_reg mstatus_ref \ ((\ (w__4 :: Mstatus) . (let p__317 = (architecture ((get_mstatus_SXL w__3 :: 2 Word.word)), (get_Mstatus_TVM w__4 :: 1 Word.word)) in (case p__317 of (Some (g__316), b__0) => if (((b__0 = ( 0b1 :: 1 Word.word)))) then handle_illegal () \ return RETIRE_FAIL else if (((b__0 = ( 0b0 :: 1 Word.word)))) then flush_TLB asid addr \ return RETIRE_SUCCESS else (case (Some g__316, b__0) of (_, _) => internal_error (''unimplemented sfence architecture'') ) | (_, _) => internal_error (''unimplemented sfence architecture'') )))))) | Machine => flush_TLB asid addr \ return RETIRE_SUCCESS ))))))))\ for rs1 :: "(5)Word.word " and rs2 :: "(5)Word.word " \ \\val execute_RTYPEW : mword ty5 -> mword ty5 -> mword ty5 -> ropw -> M Retired\\ definition execute_RTYPEW :: \(5)Word.word \(5)Word.word \(5)Word.word \ ropw \((register_value),(Retired),(exception))monad \ where \ execute_RTYPEW rs2 rs1 rd op1 = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val = ((subrange_vec_dec w__0 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val = ((subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (let (result :: 32 bits) = ((case op1 of RISCV_ADDW => (add_vec rs1_val rs2_val :: 32 Word.word) | RISCV_SUBW => (sub_vec rs1_val rs2_val :: 32 Word.word) | RISCV_SLLW => (shift_bits_left rs1_val ((subrange_vec_dec rs2_val (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) | RISCV_SRLW => (shift_bits_right rs1_val ((subrange_vec_dec rs2_val (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) | RISCV_SRAW => (shift_right_arith32 rs1_val ((subrange_vec_dec rs2_val (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) )) in wX_bits rd ((EXTS (( 32 :: int)::ii) result :: 32 Word.word)) \ return RETIRE_SUCCESS))))))))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and op1 :: " ropw " \ \\val execute_RTYPE : mword ty5 -> mword ty5 -> mword ty5 -> rop -> M Retired\\ definition execute_RTYPE :: \(5)Word.word \(5)Word.word \(5)Word.word \ rop \((register_value),(Retired),(exception))monad \ where \ execute_RTYPE rs2 rs1 rd op1 = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val . (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ rs2_val . (let (result :: xlenbits) = ((case op1 of RISCV_ADD => (add_vec rs1_val rs2_val :: 32 Word.word) | RISCV_SLT => (EXTZ (( 32 :: int)::ii) ((bool_to_bits ((zopz0zI_s rs1_val rs2_val)) :: 1 Word.word)) :: 32 Word.word) | RISCV_SLTU => (EXTZ (( 32 :: int)::ii) ((bool_to_bits ((zopz0zI_u rs1_val rs2_val)) :: 1 Word.word)) :: 32 Word.word) | RISCV_AND => (and_vec rs1_val rs2_val :: 32 Word.word) | RISCV_OR => (or_vec rs1_val rs2_val :: 32 Word.word) | RISCV_XOR => (xor_vec rs1_val rs2_val :: 32 Word.word) | RISCV_SLL => (shift_bits_left rs1_val ((subrange_vec_dec rs2_val (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) | RISCV_SRL => (shift_bits_right rs1_val ((subrange_vec_dec rs2_val (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) | RISCV_SUB => (sub_vec rs1_val rs2_val :: 32 Word.word) | RISCV_SRA => (shift_right_arith32 rs1_val ((subrange_vec_dec rs2_val (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) :: 32 Word.word) )) in wX_bits rd result \ return RETIRE_SUCCESS))))))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and op1 :: " rop " \ \\val execute_RISCV_JALR : mword ty12 -> mword ty5 -> mword ty5 -> M Retired\\ definition execute_RISCV_JALR :: \(12)Word.word \(5)Word.word \(5)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_RISCV_JALR imm rs1 rd = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let (t :: xlenbits) = ((add_vec w__0 ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) :: 32 Word.word)) in (case ((ext_control_check_addr t)) of Ext_ControlAddr_Error (e) => (let (_ :: unit) = (ext_handle_control_check_error e) in return RETIRE_FAIL) | Ext_ControlAddr_OK (addr) => (let target = ((update_vec_dec addr (( 0 :: int)::ii) B0 :: 32 Word.word)) in and_boolM (return ((bit_to_bool ((access_vec_dec target (( 1 :: int)::ii)))))) (haveRVC () \ ((\ (w__1 :: bool) . return ((\ w__1))))) \ ((\ (w__2 :: bool) . if w__2 then handle_mem_exception target (E_Fetch_Addr_Align () ) \ return RETIRE_FAIL else (get_next_pc () :: ( 32 Word.word) M) \ ((\ (w__3 :: 32 Word.word) . (wX_bits rd w__3 \ set_next_pc target) \ return RETIRE_SUCCESS))))) )))))\ for imm :: "(12)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " \ \\val execute_RISCV_JAL : mword ty21 -> mword ty5 -> M Retired\\ definition execute_RISCV_JAL :: \(21)Word.word \(5)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_RISCV_JAL imm rd = ( (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let (t :: xlenbits) = ((add_vec w__0 ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) :: 32 Word.word)) in (case ((ext_control_check_pc t)) of Ext_ControlAddr_Error (e) => (let (_ :: unit) = (ext_handle_control_check_error e) in return RETIRE_FAIL) | Ext_ControlAddr_OK (target) => and_boolM (return ((bit_to_bool ((access_vec_dec target (( 1 :: int)::ii)))))) (haveRVC () \ ((\ (w__1 :: bool) . return ((\ w__1))))) \ ((\ (w__2 :: bool) . if w__2 then handle_mem_exception target (E_Fetch_Addr_Align () ) \ return RETIRE_FAIL else (get_next_pc () :: ( 32 Word.word) M) \ ((\ (w__3 :: 32 Word.word) . (wX_bits rd w__3 \ set_next_pc target) \ return RETIRE_SUCCESS)))) )))))\ for imm :: "(21)Word.word " and rd :: "(5)Word.word " \ \\val execute_REMW : mword ty5 -> mword ty5 -> mword ty5 -> bool -> M Retired\\ definition execute_REMW :: \(5)Word.word \(5)Word.word \(5)Word.word \ bool \((register_value),(Retired),(exception))monad \ where \ execute_REMW rs2 rs1 rd s = ( haveMulDiv () \ ((\ (w__0 :: bool) . if w__0 then (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs1_val = ((subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . (let rs2_val = ((subrange_vec_dec w__2 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (let (rs1_int :: ii) = (if s then Word.sint rs1_val else Word.uint rs1_val) in (let (rs2_int :: ii) = (if s then Word.sint rs2_val else Word.uint rs2_val) in (let (r :: ii) = (if (((rs2_int = (( 0 :: int)::ii)))) then rs1_int else hardware_mod rs1_int rs2_int) in wX_bits rd ((EXTS (( 32 :: int)::ii) ((to_bits (( 32 :: int)::ii) r :: 32 Word.word)) :: 32 Word.word)) \ return RETIRE_SUCCESS))))))))) else handle_illegal () \ return RETIRE_FAIL)))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and s :: " bool " \ \\val execute_REM : mword ty5 -> mword ty5 -> mword ty5 -> bool -> M Retired\\ definition execute_REM :: \(5)Word.word \(5)Word.word \(5)Word.word \ bool \((register_value),(Retired),(exception))monad \ where \ execute_REM rs2 rs1 rd s = ( haveMulDiv () \ ((\ (w__0 :: bool) . if w__0 then (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val . (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ rs2_val . (let (rs1_int :: ii) = (if s then Word.sint rs1_val else Word.uint rs1_val) in (let (rs2_int :: ii) = (if s then Word.sint rs2_val else Word.uint rs2_val) in (let (r :: ii) = (if (((rs2_int = (( 0 :: int)::ii)))) then rs1_int else hardware_mod rs1_int rs2_int) in wX_bits rd ((to_bits (( 32 :: int)::ii) r :: 32 Word.word)) \ return RETIRE_SUCCESS))))))) else handle_illegal () \ return RETIRE_FAIL)))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and s :: " bool " \ \\val execute_MULW : mword ty5 -> mword ty5 -> mword ty5 -> M Retired\\ definition execute_MULW :: \(5)Word.word \(5)Word.word \(5)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_MULW rs2 rs1 rd = ( haveMulDiv () \ ((\ (w__0 :: bool) . if w__0 then (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs1_val = ((subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . (let rs2_val = ((subrange_vec_dec w__2 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (let (rs1_int :: ii) = (Word.sint rs1_val) in (let (rs2_int :: ii) = (Word.sint rs2_val) in (let result32 = ((subrange_vec_dec ((to_bits (( 64 :: int)::ii) ((rs1_int * rs2_int)) :: 64 Word.word)) (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (let (result :: xlenbits) = ((EXTS (( 32 :: int)::ii) result32 :: 32 Word.word)) in wX_bits rd result \ return RETIRE_SUCCESS)))))))))) else handle_illegal () \ return RETIRE_FAIL)))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " \ \\val execute_MUL : mword ty5 -> mword ty5 -> mword ty5 -> bool -> bool -> bool -> M Retired\\ definition execute_MUL :: \(5)Word.word \(5)Word.word \(5)Word.word \ bool \ bool \ bool \((register_value),(Retired),(exception))monad \ where \ execute_MUL rs2 rs1 rd high signed1 signed2 = ( haveMulDiv () \ ((\ (w__0 :: bool) . if w__0 then (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val . (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ rs2_val . (let (rs1_int :: ii) = (if signed1 then Word.sint rs1_val else Word.uint rs1_val) in (let (rs2_int :: ii) = (if signed2 then Word.sint rs2_val else Word.uint rs2_val) in (let result_wide = ((to_bits (((( 2 :: int)::ii) * (( 32 :: int)::ii))) ((rs1_int * rs2_int)) :: 64 Word.word)) in (let result = (if high then (subrange_vec_dec result_wide (((((( 2 :: int)::ii) * (( 32 :: int)::ii))) - (( 1 :: int)::ii))) (( 32 :: int)::ii) :: 32 Word.word) else (subrange_vec_dec result_wide (((( 32 :: int)::ii) - (( 1 :: int)::ii))) (( 0 :: int)::ii) :: 32 Word.word)) in wX_bits rd result \ return RETIRE_SUCCESS)))))))) else handle_illegal () \ return RETIRE_FAIL)))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and high :: " bool " and signed1 :: " bool " and signed2 :: " bool " \ \\val execute_MRET : unit -> M Retired\\ definition execute_MRET :: \ unit \((register_value),(Retired),(exception))monad \ where \ execute_MRET _ = ( read_reg cur_privilege_ref \ ((\ (w__0 :: Privilege) . if (((w__0 \ Machine))) then handle_illegal () \ return RETIRE_FAIL else if ((\ ((ext_check_xret_priv Machine)))) then (let (_ :: unit) = (ext_fail_xret_priv () ) in return RETIRE_FAIL) else read_reg cur_privilege_ref \ ((\ (w__1 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . (exception_handler w__1 (CTL_MRET () ) w__2 :: ( 32 Word.word) M) \ ((\ (w__3 :: 32 Word.word) . set_next_pc w__3 \ return RETIRE_SUCCESS)))))))))\ \ \\val execute_LOAD_FP : mword ty12 -> mword ty5 -> mword ty5 -> word_width -> M Retired\\ definition execute_LOAD_FP :: \(12)Word.word \(5)Word.word \(5)Word.word \ word_width \((register_value),(Retired),(exception))monad \ where \ execute_LOAD_FP imm rs1 rd width = ( (let (offset :: xlenbits) = ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) in ext_data_get_addr rs1 offset (Read Data) width \ ((\ (w__0 :: unit Ext_DataAddr_Check) . (case w__0 of Ext_DataAddr_Error (e) => (let (_ :: unit) = (ext_handle_data_check_error e) in return RETIRE_FAIL) | Ext_DataAddr_OK (vaddr) => if ((check_misaligned vaddr width)) then handle_mem_exception vaddr (E_Load_Addr_Align () ) \ return RETIRE_FAIL else (translateAddr vaddr (Read Data) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__1 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__1 of TR_Failure ((e, _)) => handle_mem_exception vaddr e \ return RETIRE_FAIL | TR_Address ((addr, _)) => (let (aq, rl, res) = (False, False, False) in (case (width, (( 32 :: int)::ii)) of (BYTE, g__318) => handle_illegal () \ return RETIRE_FAIL | (HALF, g__319) => handle_illegal () \ return RETIRE_FAIL | (WORD, g__320) => (mem_read (Read Data) addr (( 4 :: int)::ii) aq rl res :: ( ( 32 Word.word)MemoryOpResult) M) \ ((\ (w__2 :: ( 32 Word.word) MemoryOpResult) . process_fload32 rd vaddr w__2)) | _ => assert_exp False (''Pattern match failure at model/riscv_insts_fext.sail 331:10 - 338:11'') \ exit0 () )) ))) )))))\ for imm :: "(12)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and width :: " word_width " \ \\val execute_LOADRES : bool -> bool -> mword ty5 -> word_width -> mword ty5 -> M Retired\\ definition execute_LOADRES :: \ bool \ bool \(5)Word.word \ word_width \(5)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_LOADRES aq rl rs1 width rd = ( haveAtomics () \ ((\ (w__0 :: bool) . if w__0 then ext_data_get_addr rs1 ((zeros_implicit (( 32 :: int)::ii) :: 32 Word.word)) (Read Data) width \ ((\ (w__1 :: unit Ext_DataAddr_Check) . (case w__1 of Ext_DataAddr_Error (e) => (let (_ :: unit) = (ext_handle_data_check_error e) in return RETIRE_FAIL) | Ext_DataAddr_OK (vaddr) => (let (aligned :: bool) = ((case width of BYTE => True | HALF => (((subrange_vec_dec vaddr (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)) | WORD => (((subrange_vec_dec vaddr (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)) | DOUBLE => (((subrange_vec_dec vaddr (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)) )) in if ((\ aligned)) then handle_mem_exception vaddr (E_Load_Addr_Align () ) \ return RETIRE_FAIL else (translateAddr vaddr (Read Data) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__2 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__2 of TR_Failure ((e, _)) => handle_mem_exception vaddr e \ return RETIRE_FAIL | TR_Address ((addr, _)) => (case (width, (( 32 :: int)::ii)) of (BYTE, _) => (mem_read (Read Data) addr (( 1 :: int)::ii) aq (((aq \ rl))) True :: ( ( 8 Word.word)MemoryOpResult) M) \ ((\ (w__3 :: ( 8 Word.word) MemoryOpResult) . process_loadres rd vaddr w__3 False)) | (HALF, _) => (mem_read (Read Data) addr (( 2 :: int)::ii) aq (((aq \ rl))) True :: ( ( 16 Word.word)MemoryOpResult) M) \ ((\ (w__5 :: ( 16 Word.word) MemoryOpResult) . process_loadres rd vaddr w__5 False)) | (WORD, _) => (mem_read (Read Data) addr (( 4 :: int)::ii) aq (((aq \ rl))) True :: ( ( 32 Word.word)MemoryOpResult) M) \ ((\ (w__7 :: ( 32 Word.word) MemoryOpResult) . process_loadres rd vaddr w__7 False)) | _ => internal_error (''Unexpected AMO width'') ) )))) ))) else handle_illegal () \ return RETIRE_FAIL)))\ for aq :: " bool " and rl :: " bool " and rs1 :: "(5)Word.word " and width :: " word_width " and rd :: "(5)Word.word " \ \\val execute_LOAD : mword ty12 -> mword ty5 -> mword ty5 -> bool -> word_width -> bool -> bool -> M Retired\\ definition execute_LOAD :: \(12)Word.word \(5)Word.word \(5)Word.word \ bool \ word_width \ bool \ bool \((register_value),(Retired),(exception))monad \ where \ execute_LOAD imm rs1 rd is_unsigned width aq rl = ( (let (offset :: xlenbits) = ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) in ext_data_get_addr rs1 offset (Read Data) width \ ((\ (w__0 :: unit Ext_DataAddr_Check) . (case w__0 of Ext_DataAddr_Error (e) => (let (_ :: unit) = (ext_handle_data_check_error e) in return RETIRE_FAIL) | Ext_DataAddr_OK (vaddr) => if ((check_misaligned vaddr width)) then handle_mem_exception vaddr (E_Load_Addr_Align () ) \ return RETIRE_FAIL else (translateAddr vaddr (Read Data) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__1 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__1 of TR_Failure ((e, _)) => handle_mem_exception vaddr e \ return RETIRE_FAIL | TR_Address ((paddr, _)) => (case (width, (( 32 :: int)::ii)) of (BYTE, g__310) => (mem_read (Read Data) paddr (( 1 :: int)::ii) aq rl False :: ( ( 8 Word.word)MemoryOpResult) M) \ ((\ (w__2 :: ( 8 Word.word) MemoryOpResult) . process_load rd vaddr w__2 is_unsigned)) | (HALF, g__311) => (mem_read (Read Data) paddr (( 2 :: int)::ii) aq rl False :: ( ( 16 Word.word)MemoryOpResult) M) \ ((\ (w__4 :: ( 16 Word.word) MemoryOpResult) . process_load rd vaddr w__4 is_unsigned)) | (WORD, g__312) => (mem_read (Read Data) paddr (( 4 :: int)::ii) aq rl False :: ( ( 32 Word.word)MemoryOpResult) M) \ ((\ (w__6 :: ( 32 Word.word) MemoryOpResult) . process_load rd vaddr w__6 is_unsigned)) | _ => assert_exp False (''Pattern match failure at model/riscv_insts_base.sail 329:10 - 338:11'') \ exit0 () ) ))) )))))\ for imm :: "(12)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and is_unsigned :: " bool " and width :: " word_width " and aq :: " bool " and rl :: " bool " \ \\val execute_ITYPE : mword ty12 -> mword ty5 -> mword ty5 -> iop -> M Retired\\ definition execute_ITYPE :: \(12)Word.word \(5)Word.word \(5)Word.word \ iop \((register_value),(Retired),(exception))monad \ where \ execute_ITYPE imm rs1 rd op1 = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val . (let (immext :: xlenbits) = ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) in (let (result :: xlenbits) = ((case op1 of RISCV_ADDI => (add_vec rs1_val immext :: 32 Word.word) | RISCV_SLTI => (EXTZ (( 32 :: int)::ii) ((bool_to_bits ((zopz0zI_s rs1_val immext)) :: 1 Word.word)) :: 32 Word.word) | RISCV_SLTIU => (EXTZ (( 32 :: int)::ii) ((bool_to_bits ((zopz0zI_u rs1_val immext)) :: 1 Word.word)) :: 32 Word.word) | RISCV_ANDI => (and_vec rs1_val immext :: 32 Word.word) | RISCV_ORI => (or_vec rs1_val immext :: 32 Word.word) | RISCV_XORI => (xor_vec rs1_val immext :: 32 Word.word) )) in wX_bits rd result \ return RETIRE_SUCCESS)))))\ for imm :: "(12)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and op1 :: " iop " \ \\val execute_ILLEGAL : mword ty32 -> M Retired\\ definition execute_ILLEGAL :: \(32)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_ILLEGAL s = ( handle_illegal () \ return RETIRE_FAIL )\ for s :: "(32)Word.word " \ \\val execute_F_UN_TYPE_S : mword ty5 -> mword ty5 -> f_un_op_S -> M Retired\\ definition execute_F_UN_TYPE_S :: \(5)Word.word \(5)Word.word \ f_un_op_S \((register_value),(Retired),(exception))monad \ where \ execute_F_UN_TYPE_S arg0 arg1 arg2 = ( (let merge_var = (arg0, arg1, arg2) in (case merge_var of (rs1, rd, FCLASS_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (let (rd_val_10b :: 10 bits) = (if ((f_is_neg_inf_S rs1_val_S)) then ( 0b0000000001 :: 10 Word.word) else if ((f_is_neg_norm_S rs1_val_S)) then ( 0b0000000010 :: 10 Word.word) else if ((f_is_neg_subnorm_S rs1_val_S)) then ( 0b0000000100 :: 10 Word.word) else if ((f_is_neg_zero_S rs1_val_S)) then ( 0b0000001000 :: 10 Word.word) else if ((f_is_pos_zero_S rs1_val_S)) then ( 0b0000010000 :: 10 Word.word) else if ((f_is_pos_subnorm_S rs1_val_S)) then ( 0b0000100000 :: 10 Word.word) else if ((f_is_pos_norm_S rs1_val_S)) then ( 0b0001000000 :: 10 Word.word) else if ((f_is_pos_inf_S rs1_val_S)) then ( 0b0010000000 :: 10 Word.word) else if ((f_is_SNaN_S rs1_val_S)) then ( 0b0100000000 :: 10 Word.word) else if ((f_is_QNaN_S rs1_val_S)) then ( 0b1000000000 :: 10 Word.word) else (zeros_implicit (( 10 :: int)::ii) :: 10 Word.word)) in wX_bits rd ((EXTZ (( 32 :: int)::ii) rd_val_10b :: 32 Word.word)) \ return RETIRE_SUCCESS)))) | (rs1, rd, FMV_X_W) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((subrange_vec_dec w__0 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (let (rd_val_X :: xlenbits) = ((EXTS (( 32 :: int)::ii) rs1_val_S :: 32 Word.word)) in wX_bits rd rd_val_X \ return RETIRE_SUCCESS)))) | (rs1, rd, FMV_W_X) => (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val_X . (let rd_val_S = ((subrange_vec_dec rs1_val_X (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in wF_bits rd ((nan_box rd_val_S :: 32 Word.word)) \ return RETIRE_SUCCESS))) )))\ for arg0 :: "(5)Word.word " and arg1 :: "(5)Word.word " and arg2 :: " f_un_op_S " \ \\val execute_F_UN_RM_TYPE_S : mword ty5 -> rounding_mode -> mword ty5 -> f_un_rm_op_S -> M Retired\\ definition execute_F_UN_RM_TYPE_S :: \(5)Word.word \ rounding_mode \(5)Word.word \ f_un_rm_op_S \((register_value),(Retired),(exception))monad \ where \ execute_F_UN_RM_TYPE_S arg0 arg1 arg2 arg3 = ( (let merge_var = (arg0, arg1, arg2, arg3) in (case merge_var of (rs1, rm, rd, FSQRT_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in select_instr_or_fcsr_rm rm \ ((\ (w__1 :: rounding_mode) . (let rm_3b = ((encdec_rounding_mode_forwards w__1 :: 3 Word.word)) in (riscv_f32Sqrt rm_3b rs1_val_S :: (( 5 Word.word * 32 Word.word)) M) \ ((\ varstup . (let (fflags, rd_val_S) = varstup in (write_fflags fflags \ wF_bits rd ((nan_box rd_val_S :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs1, rm, rd, FCVT_W_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in select_instr_or_fcsr_rm rm \ ((\ (w__1 :: rounding_mode) . (let rm_3b = ((encdec_rounding_mode_forwards w__1 :: 3 Word.word)) in (riscv_f32ToI32 rm_3b rs1_val_S :: (( 5 Word.word * 32 Word.word)) M) \ ((\ varstup . (let (fflags, rd_val_W) = varstup in (write_fflags fflags \ wX_bits rd ((EXTS (( 32 :: int)::ii) rd_val_W :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs1, rm, rd, FCVT_WU_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in select_instr_or_fcsr_rm rm \ ((\ (w__1 :: rounding_mode) . (let rm_3b = ((encdec_rounding_mode_forwards w__1 :: 3 Word.word)) in (riscv_f32ToUi32 rm_3b rs1_val_S :: (( 5 Word.word * 32 Word.word)) M) \ ((\ varstup . (let (fflags, rd_val_WU) = varstup in (write_fflags fflags \ wX_bits rd ((EXTS (( 32 :: int)::ii) rd_val_WU :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs1, rm, rd, FCVT_S_W) => (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_W = ((subrange_vec_dec w__0 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in select_instr_or_fcsr_rm rm \ ((\ (w__1 :: rounding_mode) . (let rm_3b = ((encdec_rounding_mode_forwards w__1 :: 3 Word.word)) in (riscv_i32ToF32 rm_3b rs1_val_W :: (( 5 Word.word * 32 Word.word)) M) \ ((\ varstup . (let (fflags, rd_val_S) = varstup in (write_fflags fflags \ wF_bits rd ((nan_box rd_val_S :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs1, rm, rd, FCVT_S_WU) => (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_WU = ((subrange_vec_dec w__0 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in select_instr_or_fcsr_rm rm \ ((\ (w__1 :: rounding_mode) . (let rm_3b = ((encdec_rounding_mode_forwards w__1 :: 3 Word.word)) in (riscv_ui32ToF32 rm_3b rs1_val_WU :: (( 5 Word.word * 32 Word.word)) M) \ ((\ varstup . (let (fflags, rd_val_S) = varstup in (write_fflags fflags \ wF_bits rd ((nan_box rd_val_S :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs1, rm, rd, FCVT_L_S) => return RETIRE_FAIL | (rs1, rm, rd, FCVT_LU_S) => return RETIRE_FAIL | (rs1, rm, rd, FCVT_S_L) => return RETIRE_FAIL | (rs1, rm, rd, FCVT_S_LU) => return RETIRE_FAIL )))\ for arg0 :: "(5)Word.word " and arg1 :: " rounding_mode " and arg2 :: "(5)Word.word " and arg3 :: " f_un_rm_op_S " \ \\val execute_F_MADD_TYPE_S : mword ty5 -> mword ty5 -> mword ty5 -> rounding_mode -> mword ty5 -> f_madd_op_S -> M Retired\\ definition execute_F_MADD_TYPE_S :: \(5)Word.word \(5)Word.word \(5)Word.word \ rounding_mode \(5)Word.word \ f_madd_op_S \((register_value),(Retired),(exception))monad \ where \ execute_F_MADD_TYPE_S rs3 rs2 rs1 rm rd op1 = ( (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_32b = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_32b = ((nan_unbox w__1 :: 32 Word.word)) in (rF_bits rs3 :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . (let rs3_val_32b = ((nan_unbox w__2 :: 32 Word.word)) in select_instr_or_fcsr_rm rm \ ((\ (w__3 :: rounding_mode) . (let rm_3b = ((encdec_rounding_mode_forwards w__3 :: 3 Word.word)) in (case op1 of FMADD_S => (riscv_f32MulAdd rm_3b rs1_val_32b rs2_val_32b rs3_val_32b :: (( 5 Word.word * 32 Word.word)) M) | FMSUB_S => (riscv_f32MulAdd rm_3b rs1_val_32b rs2_val_32b ((negate_S rs3_val_32b :: 32 Word.word)) :: (( 5 Word.word * 32 Word.word)) M) | FNMSUB_S => (riscv_f32MulAdd rm_3b ((negate_S rs1_val_32b :: 32 Word.word)) rs2_val_32b rs3_val_32b :: (( 5 Word.word * 32 Word.word)) M) | FNMADD_S => (riscv_f32MulAdd rm_3b ((negate_S rs1_val_32b :: 32 Word.word)) rs2_val_32b ((negate_S rs3_val_32b :: 32 Word.word)) :: (( 5 Word.word * 32 Word.word)) M) ) \ ((\ varstup . (let ((fflags :: 5 bits), (rd_val_32b :: 32 bits)) = varstup in (write_fflags fflags \ wF_bits rd ((nan_box rd_val_32b :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))))))))))\ for rs3 :: "(5)Word.word " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rm :: " rounding_mode " and rd :: "(5)Word.word " and op1 :: " f_madd_op_S " \ \\val execute_F_BIN_TYPE_S : mword ty5 -> mword ty5 -> mword ty5 -> f_bin_op_S -> M Retired\\ definition execute_F_BIN_TYPE_S :: \(5)Word.word \(5)Word.word \(5)Word.word \ f_bin_op_S \((register_value),(Retired),(exception))monad \ where \ execute_F_BIN_TYPE_S arg0 arg1 arg2 arg3 = ( (let merge_var = (arg0, arg1, arg2, arg3) in (case merge_var of (rs2, rs1, rd, FSGNJ_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_S = ((nan_unbox w__1 :: 32 Word.word)) in (let (s1, e1, m1) = ((fsplit_S rs1_val_S :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let (s2, e2, m2) = ((fsplit_S rs2_val_S :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let rd_val_S = ((fmake_S s2 e1 m1 :: 32 Word.word)) in wF_bits rd ((nan_box rd_val_S :: 32 Word.word)) \ return RETIRE_SUCCESS))))))))) | (rs2, rs1, rd, FSGNJN_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_S = ((nan_unbox w__1 :: 32 Word.word)) in (let (s1, e1, m1) = ((fsplit_S rs1_val_S :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let (s2, e2, m2) = ((fsplit_S rs2_val_S :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let rd_val_S = ((fmake_S ((xor_vec ( 0b1 :: 1 Word.word) s2 :: 1 Word.word)) e1 m1 :: 32 Word.word)) in wF_bits rd ((nan_box rd_val_S :: 32 Word.word)) \ return RETIRE_SUCCESS))))))))) | (rs2, rs1, rd, FSGNJX_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_S = ((nan_unbox w__1 :: 32 Word.word)) in (let (s1, e1, m1) = ((fsplit_S rs1_val_S :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let (s2, e2, m2) = ((fsplit_S rs2_val_S :: ( 1 Word.word * 8 Word.word * 23 Word.word))) in (let rd_val_S = ((fmake_S ((xor_vec s1 s2 :: 1 Word.word)) e1 m1 :: 32 Word.word)) in wF_bits rd ((nan_box rd_val_S :: 32 Word.word)) \ return RETIRE_SUCCESS))))))))) | (rs2, rs1, rd, FMIN_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_S = ((nan_unbox w__1 :: 32 Word.word)) in (let is_quiet = True in (let (rs1_lt_rs2, fflags) = ((fle_S rs1_val_S rs2_val_S is_quiet :: (bool * 5 Word.word))) in (let rd_val_S = (if (((((f_is_NaN_S rs1_val_S)) \ ((f_is_NaN_S rs2_val_S))))) then (canonical_NaN_S () :: 32 Word.word) 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 (((((f_is_neg_zero_S rs1_val_S)) \ ((f_is_pos_zero_S rs2_val_S))))) then rs1_val_S else if (((((f_is_neg_zero_S rs2_val_S)) \ ((f_is_pos_zero_S rs1_val_S))))) then rs2_val_S else if rs1_lt_rs2 then rs1_val_S else rs2_val_S) in (accrue_fflags fflags \ wF_bits rd ((nan_box rd_val_S :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs2, rs1, rd, FMAX_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_S = ((nan_unbox w__1 :: 32 Word.word)) in (let is_quiet = True in (let (rs2_lt_rs1, fflags) = ((fle_S rs2_val_S rs1_val_S is_quiet :: (bool * 5 Word.word))) in (let rd_val_S = (if (((((f_is_NaN_S rs1_val_S)) \ ((f_is_NaN_S rs2_val_S))))) then (canonical_NaN_S () :: 32 Word.word) 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 (((((f_is_neg_zero_S rs1_val_S)) \ ((f_is_pos_zero_S rs2_val_S))))) then rs2_val_S else if (((((f_is_neg_zero_S rs2_val_S)) \ ((f_is_pos_zero_S rs1_val_S))))) then rs1_val_S else if rs2_lt_rs1 then rs1_val_S else rs2_val_S) in (accrue_fflags fflags \ wF_bits rd ((nan_box rd_val_S :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs2, rs1, rd, FEQ_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_S = ((nan_unbox w__1 :: 32 Word.word)) in (riscv_f32Eq rs1_val_S rs2_val_S :: (( 5 Word.word * 32 Word.word)) M) \ ((\ varstup . (let ((fflags :: bits_fflags), (rd_val :: bits_WU)) = varstup in (write_fflags fflags \ wX_bits rd ((EXTZ (( 32 :: int)::ii) rd_val :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs2, rs1, rd, FLT_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_S = ((nan_unbox w__1 :: 32 Word.word)) in (riscv_f32Lt rs1_val_S rs2_val_S :: (( 5 Word.word * 32 Word.word)) M) \ ((\ varstup . (let ((fflags :: bits_fflags), (rd_val :: bits_WU)) = varstup in (write_fflags fflags \ wX_bits rd ((EXTZ (( 32 :: int)::ii) rd_val :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) | (rs2, rs1, rd, FLE_S) => (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_S = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_S = ((nan_unbox w__1 :: 32 Word.word)) in (riscv_f32Le rs1_val_S rs2_val_S :: (( 5 Word.word * 32 Word.word)) M) \ ((\ varstup . (let ((fflags :: bits_fflags), (rd_val :: bits_WU)) = varstup in (write_fflags fflags \ wX_bits rd ((EXTZ (( 32 :: int)::ii) rd_val :: 32 Word.word))) \ return RETIRE_SUCCESS))))))))) )))\ for arg0 :: "(5)Word.word " and arg1 :: "(5)Word.word " and arg2 :: "(5)Word.word " and arg3 :: " f_bin_op_S " \ \\val execute_F_BIN_RM_TYPE_S : mword ty5 -> mword ty5 -> rounding_mode -> mword ty5 -> f_bin_rm_op_S -> M Retired\\ definition execute_F_BIN_RM_TYPE_S :: \(5)Word.word \(5)Word.word \ rounding_mode \(5)Word.word \ f_bin_rm_op_S \((register_value),(Retired),(exception))monad \ where \ execute_F_BIN_RM_TYPE_S rs2 rs1 rm rd op1 = ( (rF_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let rs1_val_32b = ((nan_unbox w__0 :: 32 Word.word)) in (rF_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs2_val_32b = ((nan_unbox w__1 :: 32 Word.word)) in select_instr_or_fcsr_rm rm \ ((\ (w__2 :: rounding_mode) . (let rm_3b = ((encdec_rounding_mode_forwards w__2 :: 3 Word.word)) in (case op1 of FADD_S => (riscv_f32Add rm_3b rs1_val_32b rs2_val_32b :: (( 5 Word.word * 32 Word.word)) M) | FSUB_S => (riscv_f32Sub rm_3b rs1_val_32b rs2_val_32b :: (( 5 Word.word * 32 Word.word)) M) | FMUL_S => (riscv_f32Mul rm_3b rs1_val_32b rs2_val_32b :: (( 5 Word.word * 32 Word.word)) M) | FDIV_S => (riscv_f32Div rm_3b rs1_val_32b rs2_val_32b :: (( 5 Word.word * 32 Word.word)) M) ) \ ((\ varstup . (let ((fflags :: 5 bits), (rd_val_32b :: 32 bits)) = varstup in (write_fflags fflags \ wF_bits rd ((nan_box rd_val_32b :: 32 Word.word))) \ return RETIRE_SUCCESS)))))))))))))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rm :: " rounding_mode " and rd :: "(5)Word.word " and op1 :: " f_bin_rm_op_S " \ \\val execute_FENCE_TSO : mword ty4 -> mword ty4 -> M Retired\\ definition execute_FENCE_TSO :: \(4)Word.word \(4)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_FENCE_TSO pred succ = ( (case (pred, succ) of (v__1172, v__1173) => if ((((((((subrange_vec_dec v__1172 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1173 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word))))))) then barrier (Barrier_RISCV_tso () ) else return (if ((((((((subrange_vec_dec v__1172 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1173 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word))))))) then () else (let (_ :: unit) = (print_endline (''FIXME: unsupported fence'')) in () )) ) \ return RETIRE_SUCCESS )\ for pred :: "(4)Word.word " and succ :: "(4)Word.word " \ \\val execute_FENCE_RESERVED : mword ty4 -> mword ty4 -> mword ty4 -> mword ty5 -> mword ty5 -> Retired\\ definition execute_FENCE_RESERVED :: \(4)Word.word \(4)Word.word \(4)Word.word \(5)Word.word \(5)Word.word \ Retired \ where \ execute_FENCE_RESERVED fm pred succ rs rd = ( RETIRE_SUCCESS )\ for fm :: "(4)Word.word " and pred :: "(4)Word.word " and succ :: "(4)Word.word " and rs :: "(5)Word.word " and rd :: "(5)Word.word " \ \\val execute_FENCEI_RESERVED : mword ty12 -> mword ty5 -> mword ty5 -> Retired\\ definition execute_FENCEI_RESERVED :: \(12)Word.word \(5)Word.word \(5)Word.word \ Retired \ where \ execute_FENCEI_RESERVED imm rs rd = ( RETIRE_SUCCESS )\ for imm :: "(12)Word.word " and rs :: "(5)Word.word " and rd :: "(5)Word.word " \ \\val execute_FENCEI : unit -> Retired\\ definition execute_FENCEI :: \ unit \ Retired \ where \ execute_FENCEI _ = ( RETIRE_SUCCESS )\ \ \\val execute_FENCE : mword ty4 -> mword ty4 -> M Retired\\ definition execute_FENCE :: \(4)Word.word \(4)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_FENCE pred succ = ( (case (pred, succ) of (v__1132, v__1133) => if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word))))))) then barrier (Barrier_RISCV_rw_rw () ) else if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word))))))) then barrier (Barrier_RISCV_r_rw () ) else if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word))))))) then barrier (Barrier_RISCV_r_r () ) else if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then barrier (Barrier_RISCV_rw_w () ) else if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then barrier (Barrier_RISCV_w_w () ) else if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word))))))) then barrier (Barrier_RISCV_w_rw () ) else if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word))))))) then barrier (Barrier_RISCV_rw_r () ) else if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word))))))) then barrier (Barrier_RISCV_r_w () ) else if ((((((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b01 :: 2 Word.word)))) \ (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b10 :: 2 Word.word))))))) then barrier (Barrier_RISCV_w_r () ) else return (if (((((subrange_vec_dec v__1133 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then () else if (((((subrange_vec_dec v__1132 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then () else (let (_ :: unit) = (print_endline (''FIXME: unsupported fence'')) in () )) ) \ return RETIRE_SUCCESS )\ for pred :: "(4)Word.word " and succ :: "(4)Word.word " \ \\val execute_ECALL : unit -> M Retired\\ definition execute_ECALL :: \ unit \((register_value),(Retired),(exception))monad \ where \ execute_ECALL _ = ( read_reg cur_privilege_ref \ ((\ (w__0 :: Privilege) . (let (t :: sync_exception) = ((| sync_exception_trap = ((case w__0 of User => E_U_EnvCall () | Supervisor => E_S_EnvCall () | Machine => E_M_EnvCall () )), sync_exception_excinfo = None, sync_exception_ext_exception = None |)) in read_reg cur_privilege_ref \ ((\ (w__1 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . (exception_handler w__1 (CTL_TRAP t) w__2 :: ( 32 Word.word) M) \ ((\ (w__3 :: 32 Word.word) . set_next_pc w__3 \ return RETIRE_FAIL))))))))))\ \ \\val execute_EBREAK : unit -> M Retired\\ definition execute_EBREAK :: \ unit \((register_value),(Retired),(exception))monad \ where \ execute_EBREAK _ = ( (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . handle_mem_exception w__0 (E_Breakpoint () ) \ return RETIRE_FAIL)))\ \ \\val execute_DIVW : mword ty5 -> mword ty5 -> mword ty5 -> bool -> M Retired\\ definition execute_DIVW :: \(5)Word.word \(5)Word.word \(5)Word.word \ bool \((register_value),(Retired),(exception))monad \ where \ execute_DIVW rs2 rs1 rd s = ( haveMulDiv () \ ((\ (w__0 :: bool) . if w__0 then (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (let rs1_val = ((subrange_vec_dec w__1 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__2 :: 32 Word.word) . (let rs2_val = ((subrange_vec_dec w__2 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) in (let (rs1_int :: ii) = (if s then Word.sint rs1_val else Word.uint rs1_val) in (let (rs2_int :: ii) = (if s then Word.sint rs2_val else Word.uint rs2_val) in (let (q :: ii) = (if (((rs2_int = (( 0 :: int)::ii)))) then ((( 0 :: int)-( 1 :: int))::ii) else hardware_quot rs1_int rs2_int) in (let (q' :: ii) = (if (((s \ ((q > ((((pow2 (( 31 :: int)::ii))) - (( 1 :: int)::ii)))))))) then (( 0 :: int)::ii) - ((pow (( 2 :: int)::ii) (( 31 :: int)::ii))) else q) in wX_bits rd ((EXTS (( 32 :: int)::ii) ((to_bits (( 32 :: int)::ii) q' :: 32 Word.word)) :: 32 Word.word)) \ return RETIRE_SUCCESS)))))))))) else handle_illegal () \ return RETIRE_FAIL)))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and s :: " bool " \ \\val execute_DIV : mword ty5 -> mword ty5 -> mword ty5 -> bool -> M Retired\\ definition execute_DIV :: \(5)Word.word \(5)Word.word \(5)Word.word \ bool \((register_value),(Retired),(exception))monad \ where \ execute_DIV rs2 rs1 rd s = ( haveMulDiv () \ ((\ (w__0 :: bool) . if w__0 then (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val . (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ rs2_val . (let (rs1_int :: ii) = (if s then Word.sint rs1_val else Word.uint rs1_val) in (let (rs2_int :: ii) = (if s then Word.sint rs2_val else Word.uint rs2_val) in (let (q :: ii) = (if (((rs2_int = (( 0 :: int)::ii)))) then ((( 0 :: int)-( 1 :: int))::ii) else hardware_quot rs1_int rs2_int) in (let (q' :: ii) = (if (((s \ ((q > xlen_max_signed))))) then xlen_min_signed else q) in wX_bits rd ((to_bits (( 32 :: int)::ii) q' :: 32 Word.word)) \ return RETIRE_SUCCESS)))))))) else handle_illegal () \ return RETIRE_FAIL)))\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and s :: " bool " \ \\val execute_C_SRLI_HINT : mword ty3 -> Retired\\ definition execute_C_SRLI_HINT :: \(3)Word.word \ Retired \ where \ execute_C_SRLI_HINT rsd = ( RETIRE_SUCCESS )\ for rsd :: "(3)Word.word " \ \\val execute_C_SRAI_HINT : mword ty3 -> Retired\\ definition execute_C_SRAI_HINT :: \(3)Word.word \ Retired \ where \ execute_C_SRAI_HINT rsd = ( RETIRE_SUCCESS )\ for rsd :: "(3)Word.word " \ \\val execute_C_SLLI_HINT : mword ty6 -> mword ty5 -> Retired\\ definition execute_C_SLLI_HINT :: \(6)Word.word \(5)Word.word \ Retired \ where \ execute_C_SLLI_HINT shamt rsd = ( RETIRE_SUCCESS )\ for shamt :: "(6)Word.word " and rsd :: "(5)Word.word " \ \\val execute_C_NOP_HINT : mword ty6 -> Retired\\ definition execute_C_NOP_HINT :: \(6)Word.word \ Retired \ where \ execute_C_NOP_HINT imm = ( RETIRE_SUCCESS )\ for imm :: "(6)Word.word " \ \\val execute_C_NOP : unit -> Retired\\ definition execute_C_NOP :: \ unit \ Retired \ where \ execute_C_NOP _ = ( RETIRE_SUCCESS )\ \ \\val execute_C_MV_HINT : mword ty5 -> Retired\\ definition execute_C_MV_HINT :: \(5)Word.word \ Retired \ where \ execute_C_MV_HINT rs2 = ( RETIRE_SUCCESS )\ for rs2 :: "(5)Word.word " \ \\val execute_C_LUI_HINT : mword ty6 -> Retired\\ definition execute_C_LUI_HINT :: \(6)Word.word \ Retired \ where \ execute_C_LUI_HINT imm = ( RETIRE_SUCCESS )\ for imm :: "(6)Word.word " \ \\val execute_C_LI_HINT : mword ty6 -> Retired\\ definition execute_C_LI_HINT :: \(6)Word.word \ Retired \ where \ execute_C_LI_HINT imm = ( RETIRE_SUCCESS )\ for imm :: "(6)Word.word " \ \\val execute_C_ILLEGAL : mword ty16 -> M Retired\\ definition execute_C_ILLEGAL :: \(16)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_C_ILLEGAL s = ( handle_illegal () \ return RETIRE_FAIL )\ for s :: "(16)Word.word " \ \\val execute_C_ADD_HINT : mword ty5 -> Retired\\ definition execute_C_ADD_HINT :: \(5)Word.word \ Retired \ where \ execute_C_ADD_HINT rs2 = ( RETIRE_SUCCESS )\ for rs2 :: "(5)Word.word " \ \\val execute_C_ADDI_HINT : mword ty5 -> Retired\\ definition execute_C_ADDI_HINT :: \(5)Word.word \ Retired \ where \ execute_C_ADDI_HINT rsd = ( RETIRE_SUCCESS )\ for rsd :: "(5)Word.word " \ \\val execute_CSR : mword ty12 -> mword ty5 -> mword ty5 -> bool -> csrop -> M Retired\\ definition execute_CSR :: \(12)Word.word \(5)Word.word \(5)Word.word \ bool \ csrop \((register_value),(Retired),(exception))monad \ where \ execute_CSR csr rs1 rd is_imm op1 = ( (if is_imm then return ((EXTZ (( 32 :: int)::ii) rs1 :: 32 Word.word)) else (rX_bits rs1 :: ( 32 Word.word) M)) \ ((\ (rs1_val :: xlenbits) . (let (isWrite :: bool) = ((case op1 of CSRRW => True | _ => if is_imm then (((Word.uint rs1_val)) \ (( 0 :: int)::ii)) else (((Word.uint rs1)) \ (( 0 :: int)::ii)) )) in read_reg cur_privilege_ref \ ((\ (w__1 :: Privilege) . check_CSR csr w__1 isWrite \ ((\ (w__2 :: bool) . if ((\ w__2)) then handle_illegal () \ return RETIRE_FAIL else read_reg cur_privilege_ref \ ((\ (w__3 :: Privilege) . if ((\ ((ext_check_CSR csr w__3 isWrite)))) then (let (_ :: unit) = (ext_check_CSR_fail () ) in return RETIRE_FAIL) else (readCSR csr :: ( 32 Word.word) M) \ ((\ csr_val . ((if isWrite then (let (new_val :: xlenbits) = ((case op1 of CSRRW => rs1_val | CSRRS => (or_vec csr_val rs1_val :: 32 Word.word) | CSRRC => (and_vec csr_val ((not_vec rs1_val :: 32 Word.word)) :: 32 Word.word) )) in writeCSR csr new_val) else return () ) \ wX_bits rd csr_val) \ return RETIRE_SUCCESS))))))))))))\ for csr :: "(12)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and is_imm :: " bool " and op1 :: " csrop " \ \\val execute_BTYPE : mword ty13 -> mword ty5 -> mword ty5 -> bop -> M Retired\\ definition execute_BTYPE :: \(13)Word.word \(5)Word.word \(5)Word.word \ bop \((register_value),(Retired),(exception))monad \ where \ execute_BTYPE imm rs2 rs1 op1 = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ rs1_val . (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ rs2_val . (let (taken :: bool) = ((case op1 of RISCV_BEQ => (rs1_val = rs2_val) | RISCV_BNE => (rs1_val \ rs2_val) | RISCV_BLT => zopz0zI_s rs1_val rs2_val | RISCV_BGE => zopz0zKzJ_s rs1_val rs2_val | RISCV_BLTU => zopz0zI_u rs1_val rs2_val | RISCV_BGEU => zopz0zKzJ_u rs1_val rs2_val )) in (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let (t :: xlenbits) = ((add_vec w__0 ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) :: 32 Word.word)) in if taken then (case ((ext_control_check_pc t)) of Ext_ControlAddr_Error (e) => (let (_ :: unit) = (ext_handle_control_check_error e) in return RETIRE_FAIL) | Ext_ControlAddr_OK (target) => and_boolM (return ((bit_to_bool ((access_vec_dec target (( 1 :: int)::ii)))))) (haveRVC () \ ((\ (w__1 :: bool) . return ((\ w__1))))) \ ((\ (w__2 :: bool) . if w__2 then handle_mem_exception target (E_Fetch_Addr_Align () ) \ return RETIRE_FAIL else set_next_pc target \ return RETIRE_SUCCESS)) ) else return RETIRE_SUCCESS)))))))))\ for imm :: "(13)Word.word " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and op1 :: " bop " \ \\val execute_AMO : amoop -> bool -> bool -> mword ty5 -> mword ty5 -> word_width -> mword ty5 -> M Retired\\ definition execute_AMO :: \ amoop \ bool \ bool \(5)Word.word \(5)Word.word \ word_width \(5)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_AMO op1 aq rl rs2 rs1 width rd = ( haveAtomics () \ ((\ (w__0 :: bool) . if w__0 then ext_data_get_addr rs1 ((zeros_implicit (( 32 :: int)::ii) :: 32 Word.word)) (ReadWrite (Data, Data)) width \ ((\ (w__1 :: unit Ext_DataAddr_Check) . (case w__1 of Ext_DataAddr_Error (e) => (let (_ :: unit) = (ext_handle_data_check_error e) in return RETIRE_FAIL) | Ext_DataAddr_OK (vaddr) => (translateAddr vaddr (ReadWrite (Data, Data)) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__2 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__2 of TR_Failure ((e, _)) => handle_mem_exception vaddr e \ return RETIRE_FAIL | TR_Address ((addr, _)) => (case (width, (( 32 :: int)::ii)) of (BYTE, _) => mem_write_ea addr (( 1 :: int)::ii) (((aq \ rl))) rl True | (HALF, _) => mem_write_ea addr (( 2 :: int)::ii) (((aq \ rl))) rl True | (WORD, _) => mem_write_ea addr (( 4 :: int)::ii) (((aq \ rl))) rl True | _ => internal_error (''Unexpected AMO width'') ) \ ((\ (eares :: unit MemoryOpResult) . (let (is_unsigned :: bool) = ((case op1 of AMOMINU => True | AMOMAXU => True | _ => False )) in (case width of BYTE => if is_unsigned then (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__7 :: 32 Word.word) . return ((EXTZ (( 32 :: int)::ii) ((subrange_vec_dec w__7 (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)) :: 32 Word.word)))) else (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__8 :: 32 Word.word) . return ((EXTS (( 32 :: int)::ii) ((subrange_vec_dec w__8 (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)) :: 32 Word.word)))) | HALF => if is_unsigned then (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__10 :: 32 Word.word) . return ((EXTZ (( 32 :: int)::ii) ((subrange_vec_dec w__10 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) :: 32 Word.word)))) else (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__11 :: 32 Word.word) . return ((EXTS (( 32 :: int)::ii) ((subrange_vec_dec w__11 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) :: 32 Word.word)))) | WORD => if is_unsigned then (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__13 :: 32 Word.word) . return ((EXTZ (( 32 :: int)::ii) ((subrange_vec_dec w__13 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) :: 32 Word.word)))) else (rX_bits rs2 :: ( 32 Word.word) M) \ ((\ (w__14 :: 32 Word.word) . return ((EXTS (( 32 :: int)::ii) ((subrange_vec_dec w__14 (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) :: 32 Word.word)))) | DOUBLE => (rX_bits rs2 :: ( 32 Word.word) M) ) \ ((\ (rs2_val :: xlenbits) . (case eares of MemException (e) => handle_mem_exception addr e \ return RETIRE_FAIL | MemValue (_) => (case (width, (( 32 :: int)::ii)) of (BYTE, _) => (mem_read (ReadWrite (Data, Data)) addr (( 1 :: int)::ii) aq (((aq \ rl))) True :: ( ( 8 Word.word)MemoryOpResult) M) \ ((\ (w__17 :: ( 8 Word.word) MemoryOpResult) . return ((extend_value is_unsigned w__17 :: ( 32 Word.word) MemoryOpResult)))) | (HALF, _) => (mem_read (ReadWrite (Data, Data)) addr (( 2 :: int)::ii) aq (((aq \ rl))) True :: ( ( 16 Word.word)MemoryOpResult) M) \ ((\ (w__18 :: ( 16 Word.word) MemoryOpResult) . return ((extend_value is_unsigned w__18 :: ( 32 Word.word) MemoryOpResult)))) | (WORD, _) => (mem_read (ReadWrite (Data, Data)) addr (( 4 :: int)::ii) aq (((aq \ rl))) True :: ( ( 32 Word.word)MemoryOpResult) M) \ ((\ (w__19 :: ( 32 Word.word) MemoryOpResult) . return ((extend_value is_unsigned w__19 :: ( 32 Word.word) MemoryOpResult)))) | _ => (internal_error (''Unexpected AMO width'') :: ( ( 32 Word.word)MemoryOpResult) M) ) \ ((\ (mval :: xlenbits MemoryOpResult) . (case mval of MemException (e) => handle_mem_exception addr e \ return RETIRE_FAIL | MemValue (loaded) => (let (result :: xlenbits) = ((case op1 of AMOSWAP => rs2_val | AMOADD => (add_vec rs2_val loaded :: 32 Word.word) | AMOXOR => (xor_vec rs2_val loaded :: 32 Word.word) | AMOAND => (and_vec rs2_val loaded :: 32 Word.word) | AMOOR => (or_vec rs2_val loaded :: 32 Word.word) | AMOMIN => (to_bits (( 32 :: int)::ii) ((min ((Word.sint rs2_val)) ((Word.sint loaded)))) :: 32 Word.word) | AMOMAX => (to_bits (( 32 :: int)::ii) ((max ((Word.sint rs2_val)) ((Word.sint loaded)))) :: 32 Word.word) | AMOMINU => (to_bits (( 32 :: int)::ii) ((min ((Word.uint rs2_val)) ((Word.uint loaded)))) :: 32 Word.word) | AMOMAXU => (to_bits (( 32 :: int)::ii) ((max ((Word.uint rs2_val)) ((Word.uint loaded)))) :: 32 Word.word) )) in (let (rval :: xlenbits) = ((case width of BYTE => (EXTS (( 32 :: int)::ii) ((subrange_vec_dec loaded (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)) :: 32 Word.word) | HALF => (EXTS (( 32 :: int)::ii) ((subrange_vec_dec loaded (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) :: 32 Word.word) | WORD => (EXTS (( 32 :: int)::ii) ((subrange_vec_dec loaded (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) :: 32 Word.word) | DOUBLE => loaded )) in (case (width, (( 32 :: int)::ii)) of (BYTE, _) => mem_write_value addr (( 1 :: int)::ii) ((subrange_vec_dec result (( 7 :: int)::ii) (( 0 :: int)::ii) :: 8 Word.word)) (((aq \ rl))) rl True | (HALF, _) => mem_write_value addr (( 2 :: int)::ii) ((subrange_vec_dec result (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) (((aq \ rl))) rl True | (WORD, _) => mem_write_value addr (( 4 :: int)::ii) ((subrange_vec_dec result (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) (((aq \ rl))) rl True | _ => internal_error (''Unexpected AMO width'') ) \ ((\ (wval :: bool MemoryOpResult) . (case wval of MemValue (True) => wX_bits rd rval \ return RETIRE_SUCCESS | MemValue (False) => internal_error (''AMO got false from mem_write_value'') | MemException (e) => handle_mem_exception addr e \ return RETIRE_FAIL ))))) ))) )))))) ))) ))) else handle_illegal () \ return RETIRE_FAIL)))\ for op1 :: " amoop " and aq :: " bool " and rl :: " bool " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and width :: " word_width " and rd :: "(5)Word.word " \ \\val execute_ADDIW : mword ty12 -> mword ty5 -> mword ty5 -> M Retired\\ definition execute_ADDIW :: \(12)Word.word \(5)Word.word \(5)Word.word \((register_value),(Retired),(exception))monad \ where \ execute_ADDIW imm rs1 rd = ( (rX_bits rs1 :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (let (result :: xlenbits) = ((add_vec ((EXTS (( 32 :: int)::ii) imm :: 32 Word.word)) w__0 :: 32 Word.word)) in wX_bits rd ((EXTS (( 32 :: int)::ii) ((subrange_vec_dec result (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) :: 32 Word.word)) \ return RETIRE_SUCCESS))))\ for imm :: "(12)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " function (sequential,domintros) execute :: \ ast \((register_value),(Retired),(exception))monad \ where \ execute (C_ADDI4SPN ((rdc, nzimm))) = ( (let (imm :: 12 bits) = ((concat_vec ( 0b00 :: 2 Word.word) ((concat_vec nzimm ( 0b00 :: 2 Word.word) :: 10 Word.word)) :: 12 Word.word)) in (let rd = ((creg2reg_idx rdc :: 5 Word.word)) in execute (ITYPE (imm, sp, rd, RISCV_ADDI)))))\ for nzimm :: "(8)Word.word " and rdc :: "(3)Word.word " |\ execute (C_LW ((uimm, rsc, rdc))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) in (let rd = ((creg2reg_idx rdc :: 5 Word.word)) in (let rs = ((creg2reg_idx rsc :: 5 Word.word)) in execute (LOAD (imm, rs, rd, False, WORD, False, False))))))\ for rsc :: "(3)Word.word " and uimm :: "(5)Word.word " and rdc :: "(3)Word.word " |\ execute (C_LD ((uimm, rsc, rdc))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b000 :: 3 Word.word) :: 8 Word.word)) :: 12 Word.word)) in (let rd = ((creg2reg_idx rdc :: 5 Word.word)) in (let rs = ((creg2reg_idx rsc :: 5 Word.word)) in execute (LOAD (imm, rs, rd, False, DOUBLE, False, False))))))\ for rsc :: "(3)Word.word " and uimm :: "(5)Word.word " and rdc :: "(3)Word.word " |\ execute (C_SW ((uimm, rsc1, rsc2))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) in (let rs1 = ((creg2reg_idx rsc1 :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rsc2 :: 5 Word.word)) in execute (STORE (imm, rs2, rs1, WORD, False, False))))))\ for rsc2 :: "(3)Word.word " and rsc1 :: "(3)Word.word " and uimm :: "(5)Word.word " |\ execute (C_SD ((uimm, rsc1, rsc2))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b000 :: 3 Word.word) :: 8 Word.word)) :: 12 Word.word)) in (let rs1 = ((creg2reg_idx rsc1 :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rsc2 :: 5 Word.word)) in execute (STORE (imm, rs2, rs1, DOUBLE, False, False))))))\ for rsc2 :: "(3)Word.word " and rsc1 :: "(3)Word.word " and uimm :: "(5)Word.word " |\ execute (C_ADDI ((nzi, rsd))) = ( (let (imm :: 12 bits) = ((EXTS (( 12 :: int)::ii) nzi :: 12 Word.word)) in execute (ITYPE (imm, rsd, rsd, RISCV_ADDI))))\ for nzi :: "(6)Word.word " and rsd :: "(5)Word.word " |\ execute (C_JAL (imm)) = ( execute (RISCV_JAL ((EXTS (( 21 :: int)::ii) ((concat_vec imm ( 0b0 :: 1 Word.word) :: 12 Word.word)) :: 21 Word.word), ra)))\ for imm :: "(11)Word.word " |\ execute (C_ADDIW ((imm, rsd))) = ( execute (ADDIW ((EXTS (( 12 :: int)::ii) imm :: 12 Word.word), rsd, rsd)))\ for rsd :: "(5)Word.word " and imm :: "(6)Word.word " |\ execute (C_LI ((imm, rd))) = ( (let (imm :: 12 bits) = ((EXTS (( 12 :: int)::ii) imm :: 12 Word.word)) in execute (ITYPE (imm, zreg, rd, RISCV_ADDI))))\ for imm :: "(6)Word.word " and rd :: "(5)Word.word " |\ execute (C_ADDI16SP (imm)) = ( (let (imm :: 12 bits) = ((EXTS (( 12 :: int)::ii) ((concat_vec imm ( 0x0 :: 4 Word.word) :: 10 Word.word)) :: 12 Word.word)) in execute (ITYPE (imm, sp, sp, RISCV_ADDI))))\ for imm :: "(6)Word.word " |\ execute (C_LUI ((imm, rd))) = ( (let (res :: 20 bits) = ((EXTS (( 20 :: int)::ii) imm :: 20 Word.word)) in execute (UTYPE (res, rd, RISCV_LUI))))\ for imm :: "(6)Word.word " and rd :: "(5)Word.word " |\ execute (C_SRLI ((shamt, rsd))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in execute (SHIFTIOP (shamt, rsd, rsd, RISCV_SRLI))))\ for rsd :: "(3)Word.word " and shamt :: "(6)Word.word " |\ execute (C_SRAI ((shamt, rsd))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in execute (SHIFTIOP (shamt, rsd, rsd, RISCV_SRAI))))\ for rsd :: "(3)Word.word " and shamt :: "(6)Word.word " |\ execute (C_ANDI ((imm, rsd))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in execute (ITYPE ((EXTS (( 12 :: int)::ii) imm :: 12 Word.word), rsd, rsd, RISCV_ANDI))))\ for rsd :: "(3)Word.word " and imm :: "(6)Word.word " |\ execute (C_SUB ((rsd, rs2))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rs2 :: 5 Word.word)) in execute (RTYPE (rs2, rsd, rsd, RISCV_SUB)))))\ for rsd :: "(3)Word.word " and rs2 :: "(3)Word.word " |\ execute (C_XOR ((rsd, rs2))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rs2 :: 5 Word.word)) in execute (RTYPE (rs2, rsd, rsd, RISCV_XOR)))))\ for rsd :: "(3)Word.word " and rs2 :: "(3)Word.word " |\ execute (C_OR ((rsd, rs2))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rs2 :: 5 Word.word)) in execute (RTYPE (rs2, rsd, rsd, RISCV_OR)))))\ for rsd :: "(3)Word.word " and rs2 :: "(3)Word.word " |\ execute (C_AND ((rsd, rs2))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rs2 :: 5 Word.word)) in execute (RTYPE (rs2, rsd, rsd, RISCV_AND)))))\ for rsd :: "(3)Word.word " and rs2 :: "(3)Word.word " |\ execute (C_SUBW ((rsd, rs2))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rs2 :: 5 Word.word)) in execute (RTYPEW (rs2, rsd, rsd, RISCV_SUBW)))))\ for rsd :: "(3)Word.word " and rs2 :: "(3)Word.word " |\ execute (C_ADDW ((rsd, rs2))) = ( (let rsd = ((creg2reg_idx rsd :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rs2 :: 5 Word.word)) in execute (RTYPEW (rs2, rsd, rsd, RISCV_ADDW)))))\ for rsd :: "(3)Word.word " and rs2 :: "(3)Word.word " |\ execute (C_J (imm)) = ( execute (RISCV_JAL ((EXTS (( 21 :: int)::ii) ((concat_vec imm ( 0b0 :: 1 Word.word) :: 12 Word.word)) :: 21 Word.word), zreg)))\ for imm :: "(11)Word.word " |\ execute (C_BEQZ ((imm, rs))) = ( execute (BTYPE ((EXTS (( 13 :: int)::ii) ((concat_vec imm ( 0b0 :: 1 Word.word) :: 9 Word.word)) :: 13 Word.word), zreg, (creg2reg_idx rs :: 5 Word.word), RISCV_BEQ)))\ for rs :: "(3)Word.word " and imm :: "(8)Word.word " |\ execute (C_BNEZ ((imm, rs))) = ( execute (BTYPE ((EXTS (( 13 :: int)::ii) ((concat_vec imm ( 0b0 :: 1 Word.word) :: 9 Word.word)) :: 13 Word.word), zreg, (creg2reg_idx rs :: 5 Word.word), RISCV_BNE)))\ for rs :: "(3)Word.word " and imm :: "(8)Word.word " |\ execute (C_SLLI ((shamt, rsd))) = ( execute (SHIFTIOP (shamt, rsd, rsd, RISCV_SLLI)))\ for rsd :: "(5)Word.word " and shamt :: "(6)Word.word " |\ execute (C_LWSP ((uimm, rd))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 8 Word.word)) :: 12 Word.word)) in execute (LOAD (imm, sp, rd, False, WORD, False, False))))\ for uimm :: "(6)Word.word " and rd :: "(5)Word.word " |\ execute (C_LDSP ((uimm, rd))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b000 :: 3 Word.word) :: 9 Word.word)) :: 12 Word.word)) in execute (LOAD (imm, sp, rd, False, DOUBLE, False, False))))\ for uimm :: "(6)Word.word " and rd :: "(5)Word.word " |\ execute (C_SWSP ((uimm, rs2))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 8 Word.word)) :: 12 Word.word)) in execute (STORE (imm, rs2, sp, WORD, False, False))))\ for uimm :: "(6)Word.word " and rs2 :: "(5)Word.word " |\ execute (C_SDSP ((uimm, rs2))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b000 :: 3 Word.word) :: 9 Word.word)) :: 12 Word.word)) in execute (STORE (imm, rs2, sp, DOUBLE, False, False))))\ for uimm :: "(6)Word.word " and rs2 :: "(5)Word.word " |\ execute (C_JR (rs1)) = ( execute (RISCV_JALR ((EXTZ (( 12 :: int)::ii) ( 0b0 :: 1 Word.word) :: 12 Word.word), rs1, zreg)))\ for rs1 :: "(5)Word.word " |\ execute (C_JALR (rs1)) = ( execute (RISCV_JALR ((EXTZ (( 12 :: int)::ii) ( 0b0 :: 1 Word.word) :: 12 Word.word), rs1, ra)))\ for rs1 :: "(5)Word.word " |\ execute (C_MV ((rd, rs2))) = ( execute (RTYPE (rs2, zreg, rd, RISCV_ADD)))\ for rs2 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (C_EBREAK (_)) = ( execute (EBREAK () ))\ |\ execute (C_ADD ((rsd, rs2))) = ( execute (RTYPE (rs2, rsd, rsd, RISCV_ADD)))\ for rsd :: "(5)Word.word " and rs2 :: "(5)Word.word " |\ execute (C_FLWSP ((imm, rd))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec imm ( 0b00 :: 2 Word.word) :: 8 Word.word)) :: 12 Word.word)) in execute (LOAD_FP (imm, sp, rd, WORD))))\ for imm :: "(6)Word.word " and rd :: "(5)Word.word " |\ execute (C_FSWSP ((uimm, rs2))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 8 Word.word)) :: 12 Word.word)) in execute (STORE_FP (imm, rs2, sp, WORD))))\ for uimm :: "(6)Word.word " and rs2 :: "(5)Word.word " |\ execute (C_FLW ((uimm, rsc, rdc))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) in (let rd = ((creg2reg_idx rdc :: 5 Word.word)) in (let rs = ((creg2reg_idx rsc :: 5 Word.word)) in execute (LOAD_FP (imm, rs, rd, WORD))))))\ for rsc :: "(3)Word.word " and uimm :: "(5)Word.word " and rdc :: "(3)Word.word " |\ execute (C_FSW ((uimm, rsc1, rsc2))) = ( (let (imm :: 12 bits) = ((EXTZ (( 12 :: int)::ii) ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 7 Word.word)) :: 12 Word.word)) in (let rs1 = ((creg2reg_idx rsc1 :: 5 Word.word)) in (let rs2 = ((creg2reg_idx rsc2 :: 5 Word.word)) in execute (STORE_FP (imm, rs2, rs1, WORD))))))\ for rsc2 :: "(3)Word.word " and rsc1 :: "(3)Word.word " and uimm :: "(5)Word.word " |\ execute (UTYPE ((imm, rd, op1))) = ( execute_UTYPE imm rd op1 )\ for op1 :: " uop " and imm :: "(20)Word.word " and rd :: "(5)Word.word " |\ execute (RISCV_JAL ((imm, rd))) = ( execute_RISCV_JAL imm rd )\ for imm :: "(21)Word.word " and rd :: "(5)Word.word " |\ execute (BTYPE ((imm, rs2, rs1, op1))) = ( execute_BTYPE imm rs2 rs1 op1 )\ for op1 :: " bop " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and imm :: "(13)Word.word " |\ execute (ITYPE ((imm, rs1, rd, op1))) = ( execute_ITYPE imm rs1 rd op1 )\ for op1 :: " iop " and rs1 :: "(5)Word.word " and imm :: "(12)Word.word " and rd :: "(5)Word.word " |\ execute (SHIFTIOP ((shamt, rs1, rd, op1))) = ( execute_SHIFTIOP shamt rs1 rd op1 )\ for op1 :: " sop " and shamt :: "(6)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (RTYPE ((rs2, rs1, rd, op1))) = ( execute_RTYPE rs2 rs1 rd op1 )\ for op1 :: " rop " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (LOAD ((imm, rs1, rd, is_unsigned, width, aq, rl))) = ( execute_LOAD imm rs1 rd is_unsigned width aq rl )\ for rs1 :: "(5)Word.word " and imm :: "(12)Word.word " and rd :: "(5)Word.word " and is_unsigned :: " bool " and aq :: " bool " and width :: " word_width " and rl :: " bool " |\ execute (STORE ((imm, rs2, rs1, width, aq, rl))) = ( execute_STORE imm rs2 rs1 width aq rl )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and imm :: "(12)Word.word " and aq :: " bool " and width :: " word_width " and rl :: " bool " |\ execute (ADDIW ((imm, rs1, rd))) = ( execute_ADDIW imm rs1 rd )\ for rs1 :: "(5)Word.word " and imm :: "(12)Word.word " and rd :: "(5)Word.word " |\ execute (SHIFTW ((shamt, rs1, rd, op1))) = ( execute_SHIFTW shamt rs1 rd op1 )\ for op1 :: " sop " and shamt :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (RTYPEW ((rs2, rs1, rd, op1))) = ( execute_RTYPEW rs2 rs1 rd op1 )\ for op1 :: " ropw " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (SHIFTIWOP ((shamt, rs1, rd, op1))) = ( execute_SHIFTIWOP shamt rs1 rd op1 )\ for op1 :: " sopw " and shamt :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (FENCE ((pred, succ))) = ( execute_FENCE pred succ )\ for pred :: "(4)Word.word " and succ :: "(4)Word.word " |\ execute (FENCE_TSO ((pred, succ))) = ( execute_FENCE_TSO pred succ )\ for pred :: "(4)Word.word " and succ :: "(4)Word.word " |\ execute (FENCEI (arg0)) = ( return ((execute_FENCEI arg0)))\ for arg0 :: " unit " |\ execute (ECALL (arg0)) = ( execute_ECALL arg0 )\ for arg0 :: " unit " |\ execute (MRET (arg0)) = ( execute_MRET arg0 )\ for arg0 :: " unit " |\ execute (SRET (arg0)) = ( execute_SRET arg0 )\ for arg0 :: " unit " |\ execute (EBREAK (arg0)) = ( execute_EBREAK arg0 )\ for arg0 :: " unit " |\ execute (WFI (arg0)) = ( execute_WFI arg0 )\ for arg0 :: " unit " |\ execute (SFENCE_VMA ((rs1, rs2))) = ( execute_SFENCE_VMA rs1 rs2 )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " |\ execute (LOADRES ((aq, rl, rs1, width, rd))) = ( execute_LOADRES aq rl rs1 width rd )\ for rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and aq :: " bool " and width :: " word_width " and rl :: " bool " |\ execute (STORECON ((aq, rl, rs2, rs1, width, rd))) = ( execute_STORECON aq rl rs2 rs1 width rd )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and aq :: " bool " and width :: " word_width " and rl :: " bool " |\ execute (AMO ((op1, aq, rl, rs2, rs1, width, rd))) = ( execute_AMO op1 aq rl rs2 rs1 width rd )\ for op1 :: " amoop " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and aq :: " bool " and width :: " word_width " and rl :: " bool " |\ execute (C_NOP (arg0)) = ( return ((execute_C_NOP arg0)))\ for arg0 :: " unit " |\ execute (MUL ((rs2, rs1, rd, high, signed1, signed2))) = ( execute_MUL rs2 rs1 rd high signed1 signed2 )\ for signed2 :: " bool " and signed1 :: " bool " and high :: " bool " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (DIV ((rs2, rs1, rd, s))) = ( execute_DIV rs2 rs1 rd s )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and s :: " bool " |\ execute (REM ((rs2, rs1, rd, s))) = ( execute_REM rs2 rs1 rd s )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and s :: " bool " |\ execute (MULW ((rs2, rs1, rd))) = ( execute_MULW rs2 rs1 rd )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (DIVW ((rs2, rs1, rd, s))) = ( execute_DIVW rs2 rs1 rd s )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and s :: " bool " |\ execute (REMW ((rs2, rs1, rd, s))) = ( execute_REMW rs2 rs1 rd s )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and s :: " bool " |\ execute (CSR ((csr, rs1, rd, is_imm, op1))) = ( execute_CSR csr rs1 rd is_imm op1 )\ for op1 :: " csrop " and is_imm :: " bool " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and csr :: "(12)Word.word " |\ execute (URET (arg0)) = ( execute_URET arg0 )\ for arg0 :: " unit " |\ execute (C_NOP_HINT (imm)) = ( return ((execute_C_NOP_HINT imm)))\ for imm :: "(6)Word.word " |\ execute (C_ADDI_HINT (rsd)) = ( return ((execute_C_ADDI_HINT rsd)))\ for rsd :: "(5)Word.word " |\ execute (C_LI_HINT (imm)) = ( return ((execute_C_LI_HINT imm)))\ for imm :: "(6)Word.word " |\ execute (C_LUI_HINT (imm)) = ( return ((execute_C_LUI_HINT imm)))\ for imm :: "(6)Word.word " |\ execute (C_MV_HINT (rs2)) = ( return ((execute_C_MV_HINT rs2)))\ for rs2 :: "(5)Word.word " |\ execute (C_ADD_HINT (rs2)) = ( return ((execute_C_ADD_HINT rs2)))\ for rs2 :: "(5)Word.word " |\ execute (C_SLLI_HINT ((shamt, rsd))) = ( return ((execute_C_SLLI_HINT shamt rsd)))\ for rsd :: "(5)Word.word " and shamt :: "(6)Word.word " |\ execute (C_SRLI_HINT (rsd)) = ( return ((execute_C_SRLI_HINT rsd)))\ for rsd :: "(3)Word.word " |\ execute (C_SRAI_HINT (rsd)) = ( return ((execute_C_SRAI_HINT rsd)))\ for rsd :: "(3)Word.word " |\ execute (FENCE_RESERVED ((fm, pred, succ, rs, rd))) = ( return ((execute_FENCE_RESERVED fm pred succ rs rd)))\ for rs :: "(5)Word.word " and fm :: "(4)Word.word " and rd :: "(5)Word.word " and pred :: "(4)Word.word " and succ :: "(4)Word.word " |\ execute (FENCEI_RESERVED ((imm, rs, rd))) = ( return ((execute_FENCEI_RESERVED imm rs rd)))\ for rs :: "(5)Word.word " and imm :: "(12)Word.word " and rd :: "(5)Word.word " |\ execute (LOAD_FP ((imm, rs1, rd, width))) = ( execute_LOAD_FP imm rs1 rd width )\ for rs1 :: "(5)Word.word " and imm :: "(12)Word.word " and rd :: "(5)Word.word " and width :: " word_width " |\ execute (STORE_FP ((imm, rs2, rs1, width))) = ( execute_STORE_FP imm rs2 rs1 width )\ for rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and imm :: "(12)Word.word " and width :: " word_width " |\ execute (F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, op1))) = ( execute_F_MADD_TYPE_S rs3 rs2 rs1 rm rd op1 )\ for op1 :: " f_madd_op_S " and rs3 :: "(5)Word.word " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and rm :: " rounding_mode " |\ execute (F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, op1))) = ( execute_F_BIN_RM_TYPE_S rs2 rs1 rm rd op1 )\ for op1 :: " f_bin_rm_op_S " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and rm :: " rounding_mode " |\ execute (F_UN_RM_TYPE_S ((rs1, rm, rd, arg3))) = ( execute_F_UN_RM_TYPE_S rs1 rm rd arg3 )\ for arg3 :: " f_un_rm_op_S " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " and rm :: " rounding_mode " |\ execute (F_BIN_TYPE_S ((rs2, rs1, rd, arg3))) = ( execute_F_BIN_TYPE_S rs2 rs1 rd arg3 )\ for arg3 :: " f_bin_op_S " and rs2 :: "(5)Word.word " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (F_UN_TYPE_S ((rs1, rd, arg2))) = ( execute_F_UN_TYPE_S rs1 rd arg2 )\ for arg2 :: " f_un_op_S " and rs1 :: "(5)Word.word " and rd :: "(5)Word.word " |\ execute (RISCV_JALR ((imm, rs1, rd))) = ( execute_RISCV_JALR imm rs1 rd )\ for rs1 :: "(5)Word.word " and imm :: "(12)Word.word " and rd :: "(5)Word.word " |\ execute (ILLEGAL (s)) = ( execute_ILLEGAL s )\ for s :: "(32)Word.word " |\ execute (C_ILLEGAL (s)) = ( execute_C_ILLEGAL s )\ for s :: "(16)Word.word " by pat_completeness (auto intro!: let_cong bind_cong MemoryOpResult.case_cong) definition assembly_forwards :: \ ast \((register_value),(string),(exception))monad \ where \ assembly_forwards ast = ( (let arg1 = ast in (case arg1 of UTYPE ((imm, rd, op1)) => reg_name_forwards rd \ ((\ (w__0 :: string) . return ((string_append ((utype_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__0 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) | RISCV_JAL ((imm, rd)) => reg_name_forwards rd \ ((\ (w__1 :: string) . return ((string_append (''jal'') ((string_append ((spc_forwards () )) ((string_append w__1 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) | RISCV_JALR ((imm, rs1, rd)) => reg_name_forwards rd \ ((\ (w__2 :: string) . reg_name_forwards rs1 \ ((\ (w__3 :: string) . return ((string_append (''jalr'') ((string_append ((spc_forwards () )) ((string_append w__2 ((string_append ((sep_forwards () )) ((string_append w__3 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))))))))) | BTYPE ((imm, rs2, rs1, op1)) => reg_name_forwards rs1 \ ((\ (w__4 :: string) . reg_name_forwards rs2 \ ((\ (w__5 :: string) . return ((string_append ((btype_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__4 ((string_append ((sep_forwards () )) ((string_append w__5 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))))))))) | ITYPE ((imm, rs1, rd, op1)) => reg_name_forwards rd \ ((\ (w__6 :: string) . reg_name_forwards rs1 \ ((\ (w__7 :: string) . return ((string_append ((itype_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__6 ((string_append ((sep_forwards () )) ((string_append w__7 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))))))))) | SHIFTIOP ((shamt, rs1, rd, op1)) => reg_name_forwards rd \ ((\ (w__8 :: string) . reg_name_forwards rs1 \ ((\ (w__9 :: string) . return ((string_append ((shiftiop_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__8 ((string_append ((sep_forwards () )) ((string_append w__9 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits shamt)) (''''))))))))))))))))))) | RTYPE ((rs2, rs1, rd, op1)) => reg_name_forwards rd \ ((\ (w__10 :: string) . reg_name_forwards rs1 \ ((\ (w__11 :: string) . reg_name_forwards rs2 \ ((\ (w__12 :: string) . return ((string_append ((rtype_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__10 ((string_append ((sep_forwards () )) ((string_append w__11 ((string_append ((sep_forwards () )) ((string_append w__12 (''''))))))))))))))))))))) | LOAD ((imm, rs1, rd, is_unsigned, size1, aq, rl)) => reg_name_forwards rd \ ((\ (w__13 :: string) . reg_name_forwards rs1 \ ((\ (w__14 :: string) . return ((string_append (''l'') ((string_append ((size_mnemonic_forwards size1)) ((string_append ((maybe_u_forwards is_unsigned)) ((string_append ((maybe_aq_forwards aq)) ((string_append ((maybe_rl_forwards rl)) ((string_append ((spc_forwards () )) ((string_append w__13 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) ((string_append ((opt_spc_forwards () )) ((string_append (''('') ((string_append ((opt_spc_forwards () )) ((string_append w__14 ((string_append ((opt_spc_forwards () )) ((string_append ('')'') (''''))))))))))))))))))))))))))))))))))) | STORE ((imm, rs2, rs1, size1, aq, rl)) => reg_name_forwards rs2 \ ((\ (w__15 :: string) . reg_name_forwards rs1 \ ((\ (w__16 :: string) . return ((string_append (''s'') ((string_append ((size_mnemonic_forwards size1)) ((string_append ((maybe_aq_forwards aq)) ((string_append ((maybe_rl_forwards rl)) ((string_append ((spc_forwards () )) ((string_append w__15 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) ((string_append ((opt_spc_forwards () )) ((string_append (''('') ((string_append ((opt_spc_forwards () )) ((string_append w__16 ((string_append ((opt_spc_forwards () )) ((string_append ('')'') (''''))))))))))))))))))))))))))))))))) | ADDIW ((imm, rs1, rd)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__17 :: string) . reg_name_forwards rs1 \ ((\ (w__18 :: string) . return ((string_append (''addiw'') ((string_append ((spc_forwards () )) ((string_append w__17 ((string_append ((sep_forwards () )) ((string_append w__18 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | SHIFTW ((shamt, rs1, rd, op1)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__21 :: string) . reg_name_forwards rs1 \ ((\ (w__22 :: string) . return ((string_append ((shiftw_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__21 ((string_append ((sep_forwards () )) ((string_append w__22 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits shamt)) (''''))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | RTYPEW ((rs2, rs1, rd, op1)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__25 :: string) . reg_name_forwards rs1 \ ((\ (w__26 :: string) . reg_name_forwards rs2 \ ((\ (w__27 :: string) . return ((string_append ((rtypew_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__25 ((string_append ((sep_forwards () )) ((string_append w__26 ((string_append ((sep_forwards () )) ((string_append w__27 (''''))))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | SHIFTIWOP ((shamt, rs1, rd, op1)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__30 :: string) . reg_name_forwards rs1 \ ((\ (w__31 :: string) . return ((string_append ((shiftiwop_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__30 ((string_append ((sep_forwards () )) ((string_append w__31 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits shamt)) (''''))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | FENCE ((pred, succ)) => fence_bits_forwards pred \ ((\ (w__34 :: string) . fence_bits_forwards succ \ ((\ (w__35 :: string) . return ((string_append (''fence'') ((string_append ((spc_forwards () )) ((string_append w__34 ((string_append ((sep_forwards () )) ((string_append w__35 (''''))))))))))))))) | FENCE_TSO ((pred, succ)) => fence_bits_forwards pred \ ((\ (w__36 :: string) . fence_bits_forwards succ \ ((\ (w__37 :: string) . return ((string_append (''fence.tso'') ((string_append ((spc_forwards () )) ((string_append w__36 ((string_append ((sep_forwards () )) ((string_append w__37 (''''))))))))))))))) | FENCEI (_) => return (''fence.i'') | ECALL (_) => return (''ecall'') | MRET (_) => return (''mret'') | SRET (_) => return (''sret'') | EBREAK (_) => return (''ebreak'') | WFI (_) => return (''wfi'') | SFENCE_VMA ((rs1, rs2)) => reg_name_forwards rs1 \ ((\ (w__38 :: string) . reg_name_forwards rs2 \ ((\ (w__39 :: string) . return ((string_append (''sfence.vma'') ((string_append ((spc_forwards () )) ((string_append w__38 ((string_append ((sep_forwards () )) ((string_append w__39 (''''))))))))))))))) | LOADRES ((aq, rl, rs1, size1, rd)) => reg_name_forwards rd \ ((\ (w__40 :: string) . reg_name_forwards rs1 \ ((\ (w__41 :: string) . return ((string_append (''lr.'') ((string_append ((size_mnemonic_forwards size1)) ((string_append ((maybe_aq_forwards aq)) ((string_append ((maybe_rl_forwards rl)) ((string_append ((spc_forwards () )) ((string_append w__40 ((string_append ((sep_forwards () )) ((string_append w__41 (''''))))))))))))))))))))) | STORECON ((aq, rl, rs2, rs1, size1, rd)) => reg_name_forwards rd \ ((\ (w__42 :: string) . reg_name_forwards rs1 \ ((\ (w__43 :: string) . reg_name_forwards rs2 \ ((\ (w__44 :: string) . return ((string_append (''sc.'') ((string_append ((size_mnemonic_forwards size1)) ((string_append ((maybe_aq_forwards aq)) ((string_append ((maybe_rl_forwards rl)) ((string_append ((spc_forwards () )) ((string_append w__42 ((string_append ((sep_forwards () )) ((string_append w__43 ((string_append ((sep_forwards () )) ((string_append w__44 (''''))))))))))))))))))))))))))) | AMO ((op1, aq, rl, rs2, rs1, width, rd)) => reg_name_forwards rd \ ((\ (w__45 :: string) . reg_name_forwards rs2 \ ((\ (w__46 :: string) . reg_name_forwards rs1 \ ((\ (w__47 :: string) . return ((string_append ((amo_mnemonic_forwards op1)) ((string_append (''.'') ((string_append ((size_mnemonic_forwards width)) ((string_append ((maybe_aq_forwards aq)) ((string_append ((maybe_rl_forwards rl)) ((string_append ((spc_forwards () )) ((string_append w__45 ((string_append ((sep_forwards () )) ((string_append w__46 ((string_append ((sep_forwards () )) ((string_append (''('') ((string_append w__47 ((string_append ('')'') (''''))))))))))))))))))))))))))))))))) | C_NOP (_) => return (''c.nop'') | C_ADDI4SPN ((rdc, nzimm)) => if (((nzimm \ ( 0x00 :: 8 Word.word)))) then creg_name_forwards rdc \ ((\ (w__48 :: string) . return ((string_append (''c.addi4spn'') ((string_append ((spc_forwards () )) ((string_append w__48 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits ((concat_vec nzimm ( 0b00 :: 2 Word.word) :: 10 Word.word)))) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LW ((uimm, rsc, rdc)) => creg_name_forwards rdc \ ((\ (w__51 :: string) . creg_name_forwards rsc \ ((\ (w__52 :: string) . return ((string_append (''c.lw'') ((string_append ((spc_forwards () )) ((string_append w__51 ((string_append ((sep_forwards () )) ((string_append w__52 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 7 Word.word)))) (''''))))))))))))))))))) | C_LD ((uimm, rsc, rdc)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then creg_name_forwards rdc \ ((\ (w__53 :: string) . creg_name_forwards rsc \ ((\ (w__54 :: string) . return ((string_append (''c.ld'') ((string_append ((spc_forwards () )) ((string_append w__53 ((string_append ((sep_forwards () )) ((string_append w__54 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits ((concat_vec uimm ( 0b000 :: 3 Word.word) :: 8 Word.word)))) (''''))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SW ((uimm, rsc1, rsc2)) => creg_name_forwards rsc1 \ ((\ (w__57 :: string) . creg_name_forwards rsc2 \ ((\ (w__58 :: string) . return ((string_append (''c.sw'') ((string_append ((spc_forwards () )) ((string_append w__57 ((string_append ((sep_forwards () )) ((string_append w__58 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 7 Word.word)))) (''''))))))))))))))))))) | C_SD ((uimm, rsc1, rsc2)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then creg_name_forwards rsc1 \ ((\ (w__59 :: string) . creg_name_forwards rsc2 \ ((\ (w__60 :: string) . return ((string_append (''c.sd'') ((string_append ((spc_forwards () )) ((string_append w__59 ((string_append ((sep_forwards () )) ((string_append w__60 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits ((concat_vec uimm ( 0b000 :: 3 Word.word) :: 8 Word.word)))) (''''))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDI ((nzi, rsd)) => if ((((((nzi \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))))) then reg_name_forwards rsd \ ((\ (w__63 :: string) . return ((string_append (''c.addi'') ((string_append ((spc_forwards () )) ((string_append w__63 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits nzi)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_JAL (imm) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then return ((string_append (''c.jal'') ((string_append ((spc_forwards () )) ((string_append ((decimal_string_of_bits ((concat_vec imm ( 0b0 :: 1 Word.word) :: 12 Word.word)))) (''''))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDIW ((imm, rsd)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rsd \ ((\ (w__68 :: string) . return ((string_append (''c.addiw'') ((string_append ((spc_forwards () )) ((string_append w__68 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LI ((imm, rd)) => if (((rd \ zreg))) then reg_name_forwards rd \ ((\ (w__71 :: string) . return ((string_append (''c.li'') ((string_append ((spc_forwards () )) ((string_append w__71 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDI16SP (imm) => if (((imm \ ( 0b000000 :: 6 Word.word)))) then return ((string_append (''c.addi16sp'') ((string_append ((spc_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LUI ((imm, rd)) => if ((((((rd \ zreg))) \ ((((((rd \ sp))) \ (((imm \ ( 0b000000 :: 6 Word.word)))))))))) then reg_name_forwards rd \ ((\ (w__76 :: string) . return ((string_append (''c.lui'') ((string_append ((spc_forwards () )) ((string_append w__76 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SRLI ((shamt, rsd)) => if (((shamt \ ( 0b000000 :: 6 Word.word)))) then creg_name_forwards rsd \ ((\ (w__79 :: string) . return ((string_append (''c.srli'') ((string_append ((spc_forwards () )) ((string_append w__79 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits shamt)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SRAI ((shamt, rsd)) => if (((shamt \ ( 0b000000 :: 6 Word.word)))) then creg_name_forwards rsd \ ((\ (w__82 :: string) . return ((string_append (''c.srai'') ((string_append ((spc_forwards () )) ((string_append w__82 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits shamt)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ANDI ((imm, rsd)) => creg_name_forwards rsd \ ((\ (w__85 :: string) . return ((string_append (''c.andi'') ((string_append ((spc_forwards () )) ((string_append w__85 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) | C_SUB ((rsd, rs2)) => creg_name_forwards rsd \ ((\ (w__86 :: string) . creg_name_forwards rs2 \ ((\ (w__87 :: string) . return ((string_append (''c.sub'') ((string_append ((spc_forwards () )) ((string_append w__86 ((string_append ((sep_forwards () )) ((string_append w__87 (''''))))))))))))))) | C_XOR ((rsd, rs2)) => creg_name_forwards rsd \ ((\ (w__88 :: string) . creg_name_forwards rs2 \ ((\ (w__89 :: string) . return ((string_append (''c.xor'') ((string_append ((spc_forwards () )) ((string_append w__88 ((string_append ((sep_forwards () )) ((string_append w__89 (''''))))))))))))))) | C_OR ((rsd, rs2)) => creg_name_forwards rsd \ ((\ (w__90 :: string) . creg_name_forwards rs2 \ ((\ (w__91 :: string) . return ((string_append (''c.or'') ((string_append ((spc_forwards () )) ((string_append w__90 ((string_append ((sep_forwards () )) ((string_append w__91 (''''))))))))))))))) | C_AND ((rsd, rs2)) => creg_name_forwards rsd \ ((\ (w__92 :: string) . creg_name_forwards rs2 \ ((\ (w__93 :: string) . return ((string_append (''c.and'') ((string_append ((spc_forwards () )) ((string_append w__92 ((string_append ((sep_forwards () )) ((string_append w__93 (''''))))))))))))))) | C_SUBW ((rsd, rs2)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then creg_name_forwards rsd \ ((\ (w__94 :: string) . creg_name_forwards rs2 \ ((\ (w__95 :: string) . return ((string_append (''c.subw'') ((string_append ((spc_forwards () )) ((string_append w__94 ((string_append ((sep_forwards () )) ((string_append w__95 (''''))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADDW ((rsd, rs2)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then creg_name_forwards rsd \ ((\ (w__98 :: string) . creg_name_forwards rs2 \ ((\ (w__99 :: string) . return ((string_append (''c.addw'') ((string_append ((spc_forwards () )) ((string_append w__98 ((string_append ((sep_forwards () )) ((string_append w__99 (''''))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_J (imm) => return ((string_append (''c.j'') ((string_append ((spc_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))) | C_BEQZ ((imm, rs)) => creg_name_forwards rs \ ((\ (w__102 :: string) . return ((string_append (''c.beqz'') ((string_append ((spc_forwards () )) ((string_append w__102 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) | C_BNEZ ((imm, rs)) => creg_name_forwards rs \ ((\ (w__103 :: string) . return ((string_append (''c.bnez'') ((string_append ((spc_forwards () )) ((string_append w__103 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) | C_SLLI ((shamt, rsd)) => if ((((((shamt \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))))) then reg_name_forwards rsd \ ((\ (w__104 :: string) . return ((string_append (''c.slli'') ((string_append ((spc_forwards () )) ((string_append w__104 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits shamt)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LWSP ((uimm, rd)) => if (((rd \ zreg))) then reg_name_forwards rd \ ((\ (w__107 :: string) . return ((string_append (''c.lwsp'') ((string_append ((spc_forwards () )) ((string_append w__107 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits uimm)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LDSP ((uimm, rd)) => if ((((((rd \ zreg))) \ ((((( 32 :: int):: ii) = (( 64 :: int):: ii))))))) then reg_name_forwards rd \ ((\ (w__110 :: string) . return ((string_append (''c.ldsp'') ((string_append ((spc_forwards () )) ((string_append w__110 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits uimm)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SWSP ((uimm, rd)) => reg_name_forwards rd \ ((\ (w__113 :: string) . return ((string_append (''c.swsp'') ((string_append ((spc_forwards () )) ((string_append w__113 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits uimm)) (''''))))))))))))) | C_SDSP ((uimm, rs2)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rs2 \ ((\ (w__114 :: string) . return ((string_append (''c.sdsp'') ((string_append ((spc_forwards () )) ((string_append w__114 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits uimm)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_JR (rs1) => if (((rs1 \ zreg))) then reg_name_forwards rs1 \ ((\ (w__117 :: string) . return ((string_append (''c.jr'') ((string_append ((spc_forwards () )) ((string_append w__117 (''''))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_JALR (rs1) => if (((rs1 \ zreg))) then reg_name_forwards rs1 \ ((\ (w__120 :: string) . return ((string_append (''c.jalr'') ((string_append ((spc_forwards () )) ((string_append w__120 (''''))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_MV ((rd, rs2)) => if ((((((rd \ zreg))) \ (((rs2 \ zreg)))))) then reg_name_forwards rd \ ((\ (w__123 :: string) . reg_name_forwards rs2 \ ((\ (w__124 :: string) . return ((string_append (''c.mv'') ((string_append ((spc_forwards () )) ((string_append w__123 ((string_append ((sep_forwards () )) ((string_append w__124 (''''))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_EBREAK (_) => return (''c.ebreak'') | C_ADD ((rsd, rs2)) => if ((((((rsd \ zreg))) \ (((rs2 \ zreg)))))) then reg_name_forwards rsd \ ((\ (w__127 :: string) . reg_name_forwards rs2 \ ((\ (w__128 :: string) . return ((string_append (''c.add'') ((string_append ((spc_forwards () )) ((string_append w__127 ((string_append ((sep_forwards () )) ((string_append w__128 (''''))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | MUL ((rs2, rs1, rd, high, signed1, signed2)) => reg_name_forwards rd \ ((\ (w__131 :: string) . reg_name_forwards rs1 \ ((\ (w__132 :: string) . reg_name_forwards rs2 \ ((\ (w__133 :: string) . return ((string_append ((mul_mnemonic_forwards (high, signed1, signed2))) ((string_append ((spc_forwards () )) ((string_append w__131 ((string_append ((sep_forwards () )) ((string_append w__132 ((string_append ((sep_forwards () )) ((string_append w__133 (''''))))))))))))))))))))) | DIV ((rs2, rs1, rd, s)) => reg_name_forwards rd \ ((\ (w__134 :: string) . reg_name_forwards rs1 \ ((\ (w__135 :: string) . reg_name_forwards rs2 \ ((\ (w__136 :: string) . return ((string_append (''div'') ((string_append ((maybe_not_u_forwards s)) ((string_append ((spc_forwards () )) ((string_append w__134 ((string_append ((sep_forwards () )) ((string_append w__135 ((string_append ((sep_forwards () )) ((string_append w__136 (''''))))))))))))))))))))))) | REM ((rs2, rs1, rd, s)) => reg_name_forwards rd \ ((\ (w__137 :: string) . reg_name_forwards rs1 \ ((\ (w__138 :: string) . reg_name_forwards rs2 \ ((\ (w__139 :: string) . return ((string_append (''rem'') ((string_append ((maybe_not_u_forwards s)) ((string_append ((spc_forwards () )) ((string_append w__137 ((string_append ((sep_forwards () )) ((string_append w__138 ((string_append ((sep_forwards () )) ((string_append w__139 (''''))))))))))))))))))))))) | MULW ((rs2, rs1, rd)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__140 :: string) . reg_name_forwards rs1 \ ((\ (w__141 :: string) . reg_name_forwards rs2 \ ((\ (w__142 :: string) . return ((string_append (''mulw'') ((string_append ((spc_forwards () )) ((string_append w__140 ((string_append ((sep_forwards () )) ((string_append w__141 ((string_append ((sep_forwards () )) ((string_append w__142 (''''))))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | DIVW ((rs2, rs1, rd, s)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__145 :: string) . reg_name_forwards rs1 \ ((\ (w__146 :: string) . reg_name_forwards rs2 \ ((\ (w__147 :: string) . return ((string_append (''div'') ((string_append ((maybe_not_u_forwards s)) ((string_append (''w'') ((string_append ((spc_forwards () )) ((string_append w__145 ((string_append ((sep_forwards () )) ((string_append w__146 ((string_append ((sep_forwards () )) ((string_append w__147 (''''))))))))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | REMW ((rs2, rs1, rd, s)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__150 :: string) . reg_name_forwards rs1 \ ((\ (w__151 :: string) . reg_name_forwards rs2 \ ((\ (w__152 :: string) . return ((string_append (''rem'') ((string_append ((maybe_not_u_forwards s)) ((string_append (''w'') ((string_append ((spc_forwards () )) ((string_append w__150 ((string_append ((sep_forwards () )) ((string_append w__151 ((string_append ((sep_forwards () )) ((string_append w__152 (''''))))))))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | CSR ((csr, rs1, rd, True, op1)) => reg_name_forwards rd \ ((\ (w__155 :: string) . return ((string_append ((csr_mnemonic_forwards op1)) ((string_append (''i'') ((string_append ((spc_forwards () )) ((string_append w__155 ((string_append ((sep_forwards () )) ((string_append ((csr_name_map_forwards csr)) ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits rs1)) (''''))))))))))))))))))) | CSR ((csr, rs1, rd, False, op1)) => reg_name_forwards rd \ ((\ (w__156 :: string) . reg_name_forwards rs1 \ ((\ (w__157 :: string) . return ((string_append ((csr_mnemonic_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__156 ((string_append ((sep_forwards () )) ((string_append ((csr_name_map_forwards csr)) ((string_append ((sep_forwards () )) ((string_append w__157 (''''))))))))))))))))))) | URET (_) => return (''uret'') | C_NOP_HINT (imm) => return ((string_append (''c.nop.hint.'') ((string_append ((decimal_string_of_bits imm)) (''''))))) | C_ADDI_HINT (rsd) => if (((rsd \ zreg))) then reg_name_forwards rsd \ ((\ (w__158 :: string) . return ((string_append (''c.addi.hint.'') ((string_append w__158 (''''))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_LI_HINT (imm) => return ((string_append (''c.li.hint.'') ((string_append ((decimal_string_of_bits imm)) (''''))))) | C_LUI_HINT (imm) => if (((imm \ ( 0b000000 :: 6 Word.word)))) then return ((string_append (''c.lui.hint.'') ((string_append ((decimal_string_of_bits imm)) (''''))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_MV_HINT (rs2) => if (((rs2 \ zreg))) then reg_name_forwards rs2 \ ((\ (w__163 :: string) . return ((string_append (''c.mv.hint.'') ((string_append w__163 (''''))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_ADD_HINT (rs2) => if (((rs2 \ zreg))) then reg_name_forwards rs2 \ ((\ (w__166 :: string) . return ((string_append (''c.add.hint.'') ((string_append w__166 (''''))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SLLI_HINT ((shamt, rsd)) => if ((((((shamt = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))))) then reg_name_forwards rsd \ ((\ (w__169 :: string) . return ((string_append (''c.slli.hint.'') ((string_append w__169 ((string_append (''.'') ((string_append ((decimal_string_of_bits shamt)) (''''))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_SRLI_HINT (rsd) => creg_name_forwards rsd \ ((\ (w__172 :: string) . return ((string_append (''c.srli.hint.'') ((string_append w__172 (''''))))))) | C_SRAI_HINT (rsd) => creg_name_forwards rsd \ ((\ (w__173 :: string) . return ((string_append (''c.srai.hint.'') ((string_append w__173 (''''))))))) | FENCE_RESERVED ((fm, pred, succ, rs, rd)) => if (((((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word)))))))))) then fence_bits_forwards pred \ ((\ (w__174 :: string) . fence_bits_forwards succ \ ((\ (w__175 :: string) . reg_name_forwards rs \ ((\ (w__176 :: string) . reg_name_forwards rd \ ((\ (w__177 :: string) . return ((string_append (''fence.reserved.'') ((string_append w__174 ((string_append (''.'') ((string_append w__175 ((string_append (''.'') ((string_append w__176 ((string_append (''.'') ((string_append w__177 ((string_append (''.'') ((string_append ((decimal_string_of_bits fm)) (''''))))))))))))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | FENCEI_RESERVED ((imm, rs, rd)) => if ((((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg))))))))) then reg_name_forwards rd \ ((\ (w__180 :: string) . reg_name_forwards rs \ ((\ (w__181 :: string) . return ((string_append (''fence.i.reserved.'') ((string_append w__180 ((string_append (''.'') ((string_append w__181 ((string_append (''.'') ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | LOAD_FP ((imm, rs1, rd, width)) => freg_name_forwards rd \ ((\ (w__184 :: string) . reg_name_forwards rs1 \ ((\ (w__185 :: string) . return ((string_append (''fl'') ((string_append ((size_mnemonic_forwards width)) ((string_append ((spc_forwards () )) ((string_append w__184 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) ((string_append ((opt_spc_forwards () )) ((string_append (''('') ((string_append ((opt_spc_forwards () )) ((string_append w__185 ((string_append ((opt_spc_forwards () )) ((string_append ('')'') (''''))))))))))))))))))))))))))))) | STORE_FP ((imm, rs2, rs1, width)) => freg_name_forwards rs2 \ ((\ (w__186 :: string) . reg_name_forwards rs1 \ ((\ (w__187 :: string) . return ((string_append (''fs'') ((string_append ((size_mnemonic_forwards width)) ((string_append ((spc_forwards () )) ((string_append w__186 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) ((string_append ((opt_spc_forwards () )) ((string_append (''('') ((string_append ((opt_spc_forwards () )) ((string_append w__187 ((string_append ((opt_spc_forwards () )) ((string_append ('')'') (''''))))))))))))))))))))))))))))) | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, op1)) => freg_name_forwards rd \ ((\ (w__188 :: string) . freg_name_forwards rs1 \ ((\ (w__189 :: string) . freg_name_forwards rs2 \ ((\ (w__190 :: string) . freg_name_forwards rs3 \ ((\ (w__191 :: string) . return ((string_append ((f_madd_type_mnemonic_S_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__188 ((string_append ((sep_forwards () )) ((string_append w__189 ((string_append ((sep_forwards () )) ((string_append w__190 ((string_append ((sep_forwards () )) ((string_append w__191 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))))))))))))))) | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, op1)) => freg_name_forwards rd \ ((\ (w__192 :: string) . freg_name_forwards rs1 \ ((\ (w__193 :: string) . freg_name_forwards rs2 \ ((\ (w__194 :: string) . return ((string_append ((f_bin_rm_type_mnemonic_S_forwards op1)) ((string_append ((spc_forwards () )) ((string_append w__192 ((string_append ((sep_forwards () )) ((string_append w__193 ((string_append ((sep_forwards () )) ((string_append w__194 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FSQRT_S)) => freg_name_forwards rd \ ((\ (w__195 :: string) . freg_name_forwards rs1 \ ((\ (w__196 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FSQRT_S)) ((string_append ((spc_forwards () )) ((string_append w__195 ((string_append ((sep_forwards () )) ((string_append w__196 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_W_S)) => reg_name_forwards rd \ ((\ (w__197 :: string) . freg_name_forwards rs1 \ ((\ (w__198 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FCVT_W_S)) ((string_append ((spc_forwards () )) ((string_append w__197 ((string_append ((sep_forwards () )) ((string_append w__198 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_WU_S)) => reg_name_forwards rd \ ((\ (w__199 :: string) . freg_name_forwards rs1 \ ((\ (w__200 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FCVT_WU_S)) ((string_append ((spc_forwards () )) ((string_append w__199 ((string_append ((sep_forwards () )) ((string_append w__200 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_W)) => freg_name_forwards rd \ ((\ (w__201 :: string) . reg_name_forwards rs1 \ ((\ (w__202 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FCVT_S_W)) ((string_append ((spc_forwards () )) ((string_append w__201 ((string_append ((sep_forwards () )) ((string_append w__202 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_WU)) => freg_name_forwards rd \ ((\ (w__203 :: string) . reg_name_forwards rs1 \ ((\ (w__204 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FCVT_S_WU)) ((string_append ((spc_forwards () )) ((string_append w__203 ((string_append ((sep_forwards () )) ((string_append w__204 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_L_S)) => reg_name_forwards rd \ ((\ (w__205 :: string) . freg_name_forwards rs1 \ ((\ (w__206 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FCVT_L_S)) ((string_append ((spc_forwards () )) ((string_append w__205 ((string_append ((sep_forwards () )) ((string_append w__206 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_LU_S)) => reg_name_forwards rd \ ((\ (w__207 :: string) . freg_name_forwards rs1 \ ((\ (w__208 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FCVT_LU_S)) ((string_append ((spc_forwards () )) ((string_append w__207 ((string_append ((sep_forwards () )) ((string_append w__208 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_L)) => freg_name_forwards rd \ ((\ (w__209 :: string) . reg_name_forwards rs1 \ ((\ (w__210 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FCVT_S_L)) ((string_append ((spc_forwards () )) ((string_append w__209 ((string_append ((sep_forwards () )) ((string_append w__210 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_UN_RM_TYPE_S ((rs1, rm, rd, FCVT_S_LU)) => freg_name_forwards rd \ ((\ (w__211 :: string) . reg_name_forwards rs1 \ ((\ (w__212 :: string) . return ((string_append ((f_un_rm_type_mnemonic_S_forwards FCVT_S_LU)) ((string_append ((spc_forwards () )) ((string_append w__211 ((string_append ((sep_forwards () )) ((string_append w__212 ((string_append ((sep_forwards () )) ((string_append ((frm_mnemonic_forwards rm)) (''''))))))))))))))))))) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJ_S)) => freg_name_forwards rd \ ((\ (w__213 :: string) . freg_name_forwards rs1 \ ((\ (w__214 :: string) . freg_name_forwards rs2 \ ((\ (w__215 :: string) . return ((string_append ((f_bin_type_mnemonic_S_forwards FSGNJ_S)) ((string_append ((spc_forwards () )) ((string_append w__213 ((string_append ((sep_forwards () )) ((string_append w__214 ((string_append ((sep_forwards () )) ((string_append w__215 (''''))))))))))))))))))))) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJN_S)) => freg_name_forwards rd \ ((\ (w__216 :: string) . freg_name_forwards rs1 \ ((\ (w__217 :: string) . freg_name_forwards rs2 \ ((\ (w__218 :: string) . return ((string_append ((f_bin_type_mnemonic_S_forwards FSGNJN_S)) ((string_append ((spc_forwards () )) ((string_append w__216 ((string_append ((sep_forwards () )) ((string_append w__217 ((string_append ((sep_forwards () )) ((string_append w__218 (''''))))))))))))))))))))) | F_BIN_TYPE_S ((rs2, rs1, rd, FSGNJX_S)) => freg_name_forwards rd \ ((\ (w__219 :: string) . freg_name_forwards rs1 \ ((\ (w__220 :: string) . freg_name_forwards rs2 \ ((\ (w__221 :: string) . return ((string_append ((f_bin_type_mnemonic_S_forwards FSGNJX_S)) ((string_append ((spc_forwards () )) ((string_append w__219 ((string_append ((sep_forwards () )) ((string_append w__220 ((string_append ((sep_forwards () )) ((string_append w__221 (''''))))))))))))))))))))) | F_BIN_TYPE_S ((rs2, rs1, rd, FMIN_S)) => freg_name_forwards rd \ ((\ (w__222 :: string) . freg_name_forwards rs1 \ ((\ (w__223 :: string) . freg_name_forwards rs2 \ ((\ (w__224 :: string) . return ((string_append ((f_bin_type_mnemonic_S_forwards FMIN_S)) ((string_append ((spc_forwards () )) ((string_append w__222 ((string_append ((sep_forwards () )) ((string_append w__223 ((string_append ((sep_forwards () )) ((string_append w__224 (''''))))))))))))))))))))) | F_BIN_TYPE_S ((rs2, rs1, rd, FMAX_S)) => freg_name_forwards rd \ ((\ (w__225 :: string) . freg_name_forwards rs1 \ ((\ (w__226 :: string) . freg_name_forwards rs2 \ ((\ (w__227 :: string) . return ((string_append ((f_bin_type_mnemonic_S_forwards FMAX_S)) ((string_append ((spc_forwards () )) ((string_append w__225 ((string_append ((sep_forwards () )) ((string_append w__226 ((string_append ((sep_forwards () )) ((string_append w__227 (''''))))))))))))))))))))) | F_BIN_TYPE_S ((rs2, rs1, rd, FEQ_S)) => reg_name_forwards rd \ ((\ (w__228 :: string) . freg_name_forwards rs1 \ ((\ (w__229 :: string) . freg_name_forwards rs2 \ ((\ (w__230 :: string) . return ((string_append ((f_bin_type_mnemonic_S_forwards FEQ_S)) ((string_append ((spc_forwards () )) ((string_append w__228 ((string_append ((sep_forwards () )) ((string_append w__229 ((string_append ((sep_forwards () )) ((string_append w__230 (''''))))))))))))))))))))) | F_BIN_TYPE_S ((rs2, rs1, rd, FLT_S)) => reg_name_forwards rd \ ((\ (w__231 :: string) . freg_name_forwards rs1 \ ((\ (w__232 :: string) . freg_name_forwards rs2 \ ((\ (w__233 :: string) . return ((string_append ((f_bin_type_mnemonic_S_forwards FLT_S)) ((string_append ((spc_forwards () )) ((string_append w__231 ((string_append ((sep_forwards () )) ((string_append w__232 ((string_append ((sep_forwards () )) ((string_append w__233 (''''))))))))))))))))))))) | F_BIN_TYPE_S ((rs2, rs1, rd, FLE_S)) => reg_name_forwards rd \ ((\ (w__234 :: string) . freg_name_forwards rs1 \ ((\ (w__235 :: string) . freg_name_forwards rs2 \ ((\ (w__236 :: string) . return ((string_append ((f_bin_type_mnemonic_S_forwards FLE_S)) ((string_append ((spc_forwards () )) ((string_append w__234 ((string_append ((sep_forwards () )) ((string_append w__235 ((string_append ((sep_forwards () )) ((string_append w__236 (''''))))))))))))))))))))) | F_UN_TYPE_S ((rs1, rd, FMV_X_W)) => reg_name_forwards rd \ ((\ (w__237 :: string) . freg_name_forwards rs1 \ ((\ (w__238 :: string) . return ((string_append ((f_un_type_mnemonic_S_forwards FMV_X_W)) ((string_append ((spc_forwards () )) ((string_append w__237 ((string_append ((sep_forwards () )) ((string_append w__238 (''''))))))))))))))) | F_UN_TYPE_S ((rs1, rd, FMV_W_X)) => freg_name_forwards rd \ ((\ (w__239 :: string) . reg_name_forwards rs1 \ ((\ (w__240 :: string) . return ((string_append ((f_un_type_mnemonic_S_forwards FMV_W_X)) ((string_append ((spc_forwards () )) ((string_append w__239 ((string_append ((sep_forwards () )) ((string_append w__240 (''''))))))))))))))) | F_UN_TYPE_S ((rs1, rd, FCLASS_S)) => reg_name_forwards rd \ ((\ (w__241 :: string) . freg_name_forwards rs1 \ ((\ (w__242 :: string) . return ((string_append ((f_un_type_mnemonic_S_forwards FCLASS_S)) ((string_append ((spc_forwards () )) ((string_append w__241 ((string_append ((sep_forwards () )) ((string_append w__242 (''''))))))))))))))) | C_FLWSP ((imm, rd)) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__243 :: string) . return ((string_append (''c.flwsp'') ((string_append ((spc_forwards () )) ((string_append w__243 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits imm)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_FSWSP ((uimm, rd)) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then reg_name_forwards rd \ ((\ (w__246 :: string) . return ((string_append (''c.fswsp'') ((string_append ((spc_forwards () )) ((string_append w__246 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits uimm)) (''''))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_FLW ((uimm, rsc, rdc)) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then creg_name_forwards rdc \ ((\ (w__249 :: string) . creg_name_forwards rsc \ ((\ (w__250 :: string) . return ((string_append (''c.flw'') ((string_append ((spc_forwards () )) ((string_append w__249 ((string_append ((sep_forwards () )) ((string_append w__250 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 7 Word.word)))) (''''))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | C_FSW ((uimm, rsc1, rsc2)) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then creg_name_forwards rsc1 \ ((\ (w__253 :: string) . creg_name_forwards rsc2 \ ((\ (w__254 :: string) . return ((string_append (''c.fsw'') ((string_append ((spc_forwards () )) ((string_append w__253 ((string_append ((sep_forwards () )) ((string_append w__254 ((string_append ((sep_forwards () )) ((string_append ((decimal_string_of_bits ((concat_vec uimm ( 0b00 :: 2 Word.word) :: 7 Word.word)))) (''''))))))))))))))))))) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () | ILLEGAL (s) => return ((string_append (''illegal'') ((string_append ((spc_forwards () )) ((string_append ((decimal_string_of_bits s)) (''''))))))) | C_ILLEGAL (s) => return ((string_append (''c.illegal'') ((string_append ((spc_forwards () )) ((string_append ((decimal_string_of_bits s)) (''''))))))) )))\ for ast :: " ast " \ \\val _s2645_ : string -> maybe (mword ty16)\\ definition s2645 :: \ string \((16)Word.word)option \ where \ s2645 s26460 = ( (let s26470 = s26460 in if ((string_startswith s26470 (''c.illegal''))) then (case ((string_drop s26470 ((string_length (''c.illegal''))))) of s26480 => (case ((spc_matches_prefix0 s26480)) of Some ((_, s26490)) => (case ((string_drop s26480 s26490)) of s26500 => (case ((hex_bits_16_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s26500 :: (( 16 Word.word * ii)) option)) of Some ((s, s26510)) => (let p00 = (string_drop s26500 s26510) in if (((p00 = ('''')))) then Some s else None) | _ => None ) ) | _ => None ) ) else None))\ for s26460 :: " string " \ \\val _s2637_ : string -> maybe (mword ty32)\\ definition s2637 :: \ string \((32)Word.word)option \ where \ s2637 s26380 = ( (let s26390 = s26380 in if ((string_startswith s26390 (''illegal''))) then (case ((string_drop s26390 ((string_length (''illegal''))))) of s26400 => (case ((spc_matches_prefix0 s26400)) of Some ((_, s26410)) => (case ((string_drop s26400 s26410)) of s26420 => (case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s26420 :: (( 32 Word.word * ii)) option)) of Some ((s, s26430)) => (let p00 = (string_drop s26420 s26430) in if (((p00 = ('''')))) then Some s else None) | _ => None ) ) | _ => None ) ) else None))\ for s26380 :: " string " \ \\val _s2621_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s2621 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s2621 s26220 = ( (let s26230 = s26220 in if ((string_startswith s26230 (''c.fsw''))) then (case ((string_drop s26230 ((string_length (''c.fsw''))))) of s26240 => (case ((spc_matches_prefix0 s26240)) of Some ((_, s26250)) => (case ((string_drop s26240 s26250)) of s26260 => (case ((creg_name_matches_prefix s26260 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s26270)) => (case ((string_drop s26260 s26270)) of s26280 => (case ((sep_matches_prefix s26280)) of Some ((_, s26290)) => (case ((string_drop s26280 s26290)) of s26300 => (case ((creg_name_matches_prefix s26300 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s26310)) => (case ((string_drop s26300 s26310)) of s26320 => (case ((sep_matches_prefix s26320)) of Some ((_, s26330)) => (case ((string_drop s26320 s26330)) of s26340 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s26340 :: (( 7 Word.word * ii)) option)) of Some ((v__1180, s26350)) => if (((((subrange_vec_dec v__1180 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1180 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1180 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s26340 s26350) in if (((p00 = ('''')))) then Some (rsc1, rsc2, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s26220 :: " string " \ \\val _s2605_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s2605 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s2605 s26060 = ( (let s26070 = s26060 in if ((string_startswith s26070 (''c.flw''))) then (case ((string_drop s26070 ((string_length (''c.flw''))))) of s26080 => (case ((spc_matches_prefix0 s26080)) of Some ((_, s26090)) => (case ((string_drop s26080 s26090)) of s26100 => (case ((creg_name_matches_prefix s26100 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s26110)) => (case ((string_drop s26100 s26110)) of s26120 => (case ((sep_matches_prefix s26120)) of Some ((_, s26130)) => (case ((string_drop s26120 s26130)) of s26140 => (case ((creg_name_matches_prefix s26140 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s26150)) => (case ((string_drop s26140 s26150)) of s26160 => (case ((sep_matches_prefix s26160)) of Some ((_, s26170)) => (case ((string_drop s26160 s26170)) of s26180 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s26180 :: (( 7 Word.word * ii)) option)) of Some ((v__1182, s26190)) => if (((((subrange_vec_dec v__1182 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1182 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1182 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s26180 s26190) in if (((p00 = ('''')))) then Some (rdc, rsc, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s26060 :: " string " \ \\val _s2593_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s2593 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s2593 s25940 = ( (let s25950 = s25940 in if ((string_startswith s25950 (''c.fswsp''))) then (case ((string_drop s25950 ((string_length (''c.fswsp''))))) of s25960 => (case ((spc_matches_prefix0 s25960)) of Some ((_, s25970)) => (case ((string_drop s25960 s25970)) of s25980 => (case ((reg_name_matches_prefix s25980 :: (( 5 Word.word * ii)) option)) of Some ((rd, s25990)) => (case ((string_drop s25980 s25990)) of s26000 => (case ((sep_matches_prefix s26000)) of Some ((_, s26010)) => (case ((string_drop s26000 s26010)) of s26020 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s26020 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s26030)) => (let p00 = (string_drop s26020 s26030) in if (((p00 = ('''')))) then Some (rd, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s25940 :: " string " \ \\val _s2581_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s2581 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s2581 s25820 = ( (let s25830 = s25820 in if ((string_startswith s25830 (''c.flwsp''))) then (case ((string_drop s25830 ((string_length (''c.flwsp''))))) of s25840 => (case ((spc_matches_prefix0 s25840)) of Some ((_, s25850)) => (case ((string_drop s25840 s25850)) of s25860 => (case ((reg_name_matches_prefix s25860 :: (( 5 Word.word * ii)) option)) of Some ((rd, s25870)) => (case ((string_drop s25860 s25870)) of s25880 => (case ((sep_matches_prefix s25880)) of Some ((_, s25890)) => (case ((string_drop s25880 s25890)) of s25900 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s25900 :: (( 6 Word.word * ii)) option)) of Some ((imm, s25910)) => (let p00 = (string_drop s25900 s25910) in if (((p00 = ('''')))) then Some (rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s25820 :: " string " \ \\val _s2568_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5))\\ definition s2568 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word)option \ where \ s2568 s25700 = ( (case ((f_un_type_mnemonic_S_matches_prefix s25700)) of Some ((FCLASS_S, s25710)) => (case ((string_drop s25700 s25710)) of s25720 => (case ((spc_matches_prefix0 s25720)) of Some ((_, s25730)) => (case ((string_drop s25720 s25730)) of s25740 => (case ((reg_name_matches_prefix s25740 :: (( 5 Word.word * ii)) option)) of Some ((rd, s25750)) => (case ((string_drop s25740 s25750)) of s25760 => (case ((sep_matches_prefix s25760)) of Some ((_, s25770)) => (case ((string_drop s25760 s25770)) of s25780 => (case ((freg_name_matches_prefix s25780 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s25790)) => (let p00 = (string_drop s25780 s25790) in if (((p00 = ('''')))) then Some (FCLASS_S, rd, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s25700 :: " string " \ \\val _s2555_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5))\\ definition s2555 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word)option \ where \ s2555 s25570 = ( (case ((f_un_type_mnemonic_S_matches_prefix s25570)) of Some ((FMV_W_X, s25580)) => (case ((string_drop s25570 s25580)) of s25590 => (case ((spc_matches_prefix0 s25590)) of Some ((_, s25600)) => (case ((string_drop s25590 s25600)) of s25610 => (case ((freg_name_matches_prefix s25610 :: (( 5 Word.word * ii)) option)) of Some ((rd, s25620)) => (case ((string_drop s25610 s25620)) of s25630 => (case ((sep_matches_prefix s25630)) of Some ((_, s25640)) => (case ((string_drop s25630 s25640)) of s25650 => (case ((reg_name_matches_prefix s25650 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s25660)) => (let p00 = (string_drop s25650 s25660) in if (((p00 = ('''')))) then Some (FMV_W_X, rd, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s25570 :: " string " \ \\val _s2542_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5))\\ definition s2542 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word)option \ where \ s2542 s25440 = ( (case ((f_un_type_mnemonic_S_matches_prefix s25440)) of Some ((FMV_X_W, s25450)) => (case ((string_drop s25440 s25450)) of s25460 => (case ((spc_matches_prefix0 s25460)) of Some ((_, s25470)) => (case ((string_drop s25460 s25470)) of s25480 => (case ((reg_name_matches_prefix s25480 :: (( 5 Word.word * ii)) option)) of Some ((rd, s25490)) => (case ((string_drop s25480 s25490)) of s25500 => (case ((sep_matches_prefix s25500)) of Some ((_, s25510)) => (case ((string_drop s25500 s25510)) of s25520 => (case ((freg_name_matches_prefix s25520 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s25530)) => (let p00 = (string_drop s25520 s25530) in if (((p00 = ('''')))) then Some (FMV_X_W, rd, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s25440 :: " string " \ \\val _s2525_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s2525 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2525 s25270 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s25270)) of Some ((FLE_S, s25280)) => (case ((string_drop s25270 s25280)) of s25290 => (case ((spc_matches_prefix0 s25290)) of Some ((_, s25300)) => (case ((string_drop s25290 s25300)) of s25310 => (case ((reg_name_matches_prefix s25310 :: (( 5 Word.word * ii)) option)) of Some ((rd, s25320)) => (case ((string_drop s25310 s25320)) of s25330 => (case ((sep_matches_prefix s25330)) of Some ((_, s25340)) => (case ((string_drop s25330 s25340)) of s25350 => (case ((freg_name_matches_prefix s25350 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s25360)) => (case ((string_drop s25350 s25360)) of s25370 => (case ((sep_matches_prefix s25370)) of Some ((_, s25380)) => (case ((string_drop s25370 s25380)) of s25390 => (case ((freg_name_matches_prefix s25390 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s25400)) => (let p00 = (string_drop s25390 s25400) in if (((p00 = ('''')))) then Some (FLE_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s25270 :: " string " \ \\val _s2508_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s2508 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2508 s25100 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s25100)) of Some ((FLT_S, s25110)) => (case ((string_drop s25100 s25110)) of s25120 => (case ((spc_matches_prefix0 s25120)) of Some ((_, s25130)) => (case ((string_drop s25120 s25130)) of s25140 => (case ((reg_name_matches_prefix s25140 :: (( 5 Word.word * ii)) option)) of Some ((rd, s25150)) => (case ((string_drop s25140 s25150)) of s25160 => (case ((sep_matches_prefix s25160)) of Some ((_, s25170)) => (case ((string_drop s25160 s25170)) of s25180 => (case ((freg_name_matches_prefix s25180 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s25190)) => (case ((string_drop s25180 s25190)) of s25200 => (case ((sep_matches_prefix s25200)) of Some ((_, s25210)) => (case ((string_drop s25200 s25210)) of s25220 => (case ((freg_name_matches_prefix s25220 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s25230)) => (let p00 = (string_drop s25220 s25230) in if (((p00 = ('''')))) then Some (FLT_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s25100 :: " string " \ \\val _s2491_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s2491 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2491 s24930 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s24930)) of Some ((FEQ_S, s24940)) => (case ((string_drop s24930 s24940)) of s24950 => (case ((spc_matches_prefix0 s24950)) of Some ((_, s24960)) => (case ((string_drop s24950 s24960)) of s24970 => (case ((reg_name_matches_prefix s24970 :: (( 5 Word.word * ii)) option)) of Some ((rd, s24980)) => (case ((string_drop s24970 s24980)) of s24990 => (case ((sep_matches_prefix s24990)) of Some ((_, s25000)) => (case ((string_drop s24990 s25000)) of s25010 => (case ((freg_name_matches_prefix s25010 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s25020)) => (case ((string_drop s25010 s25020)) of s25030 => (case ((sep_matches_prefix s25030)) of Some ((_, s25040)) => (case ((string_drop s25030 s25040)) of s25050 => (case ((freg_name_matches_prefix s25050 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s25060)) => (let p00 = (string_drop s25050 s25060) in if (((p00 = ('''')))) then Some (FEQ_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s24930 :: " string " \ \\val _s2474_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s2474 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2474 s24760 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s24760)) of Some ((FMAX_S, s24770)) => (case ((string_drop s24760 s24770)) of s24780 => (case ((spc_matches_prefix0 s24780)) of Some ((_, s24790)) => (case ((string_drop s24780 s24790)) of s24800 => (case ((freg_name_matches_prefix s24800 :: (( 5 Word.word * ii)) option)) of Some ((rd, s24810)) => (case ((string_drop s24800 s24810)) of s24820 => (case ((sep_matches_prefix s24820)) of Some ((_, s24830)) => (case ((string_drop s24820 s24830)) of s24840 => (case ((freg_name_matches_prefix s24840 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s24850)) => (case ((string_drop s24840 s24850)) of s24860 => (case ((sep_matches_prefix s24860)) of Some ((_, s24870)) => (case ((string_drop s24860 s24870)) of s24880 => (case ((freg_name_matches_prefix s24880 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s24890)) => (let p00 = (string_drop s24880 s24890) in if (((p00 = ('''')))) then Some (FMAX_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s24760 :: " string " \ \\val _s2457_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s2457 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2457 s24590 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s24590)) of Some ((FMIN_S, s24600)) => (case ((string_drop s24590 s24600)) of s24610 => (case ((spc_matches_prefix0 s24610)) of Some ((_, s24620)) => (case ((string_drop s24610 s24620)) of s24630 => (case ((freg_name_matches_prefix s24630 :: (( 5 Word.word * ii)) option)) of Some ((rd, s24640)) => (case ((string_drop s24630 s24640)) of s24650 => (case ((sep_matches_prefix s24650)) of Some ((_, s24660)) => (case ((string_drop s24650 s24660)) of s24670 => (case ((freg_name_matches_prefix s24670 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s24680)) => (case ((string_drop s24670 s24680)) of s24690 => (case ((sep_matches_prefix s24690)) of Some ((_, s24700)) => (case ((string_drop s24690 s24700)) of s24710 => (case ((freg_name_matches_prefix s24710 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s24720)) => (let p00 = (string_drop s24710 s24720) in if (((p00 = ('''')))) then Some (FMIN_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s24590 :: " string " \ \\val _s2440_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s2440 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2440 s24420 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s24420)) of Some ((FSGNJX_S, s24430)) => (case ((string_drop s24420 s24430)) of s24440 => (case ((spc_matches_prefix0 s24440)) of Some ((_, s24450)) => (case ((string_drop s24440 s24450)) of s24460 => (case ((freg_name_matches_prefix s24460 :: (( 5 Word.word * ii)) option)) of Some ((rd, s24470)) => (case ((string_drop s24460 s24470)) of s24480 => (case ((sep_matches_prefix s24480)) of Some ((_, s24490)) => (case ((string_drop s24480 s24490)) of s24500 => (case ((freg_name_matches_prefix s24500 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s24510)) => (case ((string_drop s24500 s24510)) of s24520 => (case ((sep_matches_prefix s24520)) of Some ((_, s24530)) => (case ((string_drop s24520 s24530)) of s24540 => (case ((freg_name_matches_prefix s24540 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s24550)) => (let p00 = (string_drop s24540 s24550) in if (((p00 = ('''')))) then Some (FSGNJX_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s24420 :: " string " \ \\val _s2423_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s2423 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2423 s24250 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s24250)) of Some ((FSGNJN_S, s24260)) => (case ((string_drop s24250 s24260)) of s24270 => (case ((spc_matches_prefix0 s24270)) of Some ((_, s24280)) => (case ((string_drop s24270 s24280)) of s24290 => (case ((freg_name_matches_prefix s24290 :: (( 5 Word.word * ii)) option)) of Some ((rd, s24300)) => (case ((string_drop s24290 s24300)) of s24310 => (case ((sep_matches_prefix s24310)) of Some ((_, s24320)) => (case ((string_drop s24310 s24320)) of s24330 => (case ((freg_name_matches_prefix s24330 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s24340)) => (case ((string_drop s24330 s24340)) of s24350 => (case ((sep_matches_prefix s24350)) of Some ((_, s24360)) => (case ((string_drop s24350 s24360)) of s24370 => (case ((freg_name_matches_prefix s24370 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s24380)) => (let p00 = (string_drop s24370 s24380) in if (((p00 = ('''')))) then Some (FSGNJN_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s24250 :: " string " \ \\val _s2406_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s2406 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2406 s24080 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s24080)) of Some ((FSGNJ_S, s24090)) => (case ((string_drop s24080 s24090)) of s24100 => (case ((spc_matches_prefix0 s24100)) of Some ((_, s24110)) => (case ((string_drop s24100 s24110)) of s24120 => (case ((freg_name_matches_prefix s24120 :: (( 5 Word.word * ii)) option)) of Some ((rd, s24130)) => (case ((string_drop s24120 s24130)) of s24140 => (case ((sep_matches_prefix s24140)) of Some ((_, s24150)) => (case ((string_drop s24140 s24150)) of s24160 => (case ((freg_name_matches_prefix s24160 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s24170)) => (case ((string_drop s24160 s24170)) of s24180 => (case ((sep_matches_prefix s24180)) of Some ((_, s24190)) => (case ((string_drop s24180 s24190)) of s24200 => (case ((freg_name_matches_prefix s24200 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s24210)) => (let p00 = (string_drop s24200 s24210) in if (((p00 = ('''')))) then Some (FSGNJ_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s24080 :: " string " \ \\val _s2389_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2389 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2389 s23910 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s23910)) of Some ((FCVT_S_LU, s23920)) => (case ((string_drop s23910 s23920)) of s23930 => (case ((spc_matches_prefix0 s23930)) of Some ((_, s23940)) => (case ((string_drop s23930 s23940)) of s23950 => (case ((freg_name_matches_prefix s23950 :: (( 5 Word.word * ii)) option)) of Some ((rd, s23960)) => (case ((string_drop s23950 s23960)) of s23970 => (case ((sep_matches_prefix s23970)) of Some ((_, s23980)) => (case ((string_drop s23970 s23980)) of s23990 => (case ((reg_name_matches_prefix s23990 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s24000)) => (case ((string_drop s23990 s24000)) of s24010 => (case ((sep_matches_prefix s24010)) of Some ((_, s24020)) => (case ((string_drop s24010 s24020)) of s24030 => (case ((frm_mnemonic_matches_prefix s24030)) of Some ((rm, s24040)) => (let p00 = (string_drop s24030 s24040) in if (((p00 = ('''')))) then Some (FCVT_S_LU, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s23910 :: " string " \ \\val _s2372_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2372 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2372 s23740 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s23740)) of Some ((FCVT_S_L, s23750)) => (case ((string_drop s23740 s23750)) of s23760 => (case ((spc_matches_prefix0 s23760)) of Some ((_, s23770)) => (case ((string_drop s23760 s23770)) of s23780 => (case ((freg_name_matches_prefix s23780 :: (( 5 Word.word * ii)) option)) of Some ((rd, s23790)) => (case ((string_drop s23780 s23790)) of s23800 => (case ((sep_matches_prefix s23800)) of Some ((_, s23810)) => (case ((string_drop s23800 s23810)) of s23820 => (case ((reg_name_matches_prefix s23820 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s23830)) => (case ((string_drop s23820 s23830)) of s23840 => (case ((sep_matches_prefix s23840)) of Some ((_, s23850)) => (case ((string_drop s23840 s23850)) of s23860 => (case ((frm_mnemonic_matches_prefix s23860)) of Some ((rm, s23870)) => (let p00 = (string_drop s23860 s23870) in if (((p00 = ('''')))) then Some (FCVT_S_L, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s23740 :: " string " \ \\val _s2355_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2355 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2355 s23570 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s23570)) of Some ((FCVT_LU_S, s23580)) => (case ((string_drop s23570 s23580)) of s23590 => (case ((spc_matches_prefix0 s23590)) of Some ((_, s23600)) => (case ((string_drop s23590 s23600)) of s23610 => (case ((reg_name_matches_prefix s23610 :: (( 5 Word.word * ii)) option)) of Some ((rd, s23620)) => (case ((string_drop s23610 s23620)) of s23630 => (case ((sep_matches_prefix s23630)) of Some ((_, s23640)) => (case ((string_drop s23630 s23640)) of s23650 => (case ((freg_name_matches_prefix s23650 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s23660)) => (case ((string_drop s23650 s23660)) of s23670 => (case ((sep_matches_prefix s23670)) of Some ((_, s23680)) => (case ((string_drop s23670 s23680)) of s23690 => (case ((frm_mnemonic_matches_prefix s23690)) of Some ((rm, s23700)) => (let p00 = (string_drop s23690 s23700) in if (((p00 = ('''')))) then Some (FCVT_LU_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s23570 :: " string " \ \\val _s2338_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2338 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2338 s23400 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s23400)) of Some ((FCVT_L_S, s23410)) => (case ((string_drop s23400 s23410)) of s23420 => (case ((spc_matches_prefix0 s23420)) of Some ((_, s23430)) => (case ((string_drop s23420 s23430)) of s23440 => (case ((reg_name_matches_prefix s23440 :: (( 5 Word.word * ii)) option)) of Some ((rd, s23450)) => (case ((string_drop s23440 s23450)) of s23460 => (case ((sep_matches_prefix s23460)) of Some ((_, s23470)) => (case ((string_drop s23460 s23470)) of s23480 => (case ((freg_name_matches_prefix s23480 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s23490)) => (case ((string_drop s23480 s23490)) of s23500 => (case ((sep_matches_prefix s23500)) of Some ((_, s23510)) => (case ((string_drop s23500 s23510)) of s23520 => (case ((frm_mnemonic_matches_prefix s23520)) of Some ((rm, s23530)) => (let p00 = (string_drop s23520 s23530) in if (((p00 = ('''')))) then Some (FCVT_L_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s23400 :: " string " \ \\val _s2321_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2321 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2321 s23230 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s23230)) of Some ((FCVT_S_WU, s23240)) => (case ((string_drop s23230 s23240)) of s23250 => (case ((spc_matches_prefix0 s23250)) of Some ((_, s23260)) => (case ((string_drop s23250 s23260)) of s23270 => (case ((freg_name_matches_prefix s23270 :: (( 5 Word.word * ii)) option)) of Some ((rd, s23280)) => (case ((string_drop s23270 s23280)) of s23290 => (case ((sep_matches_prefix s23290)) of Some ((_, s23300)) => (case ((string_drop s23290 s23300)) of s23310 => (case ((reg_name_matches_prefix s23310 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s23320)) => (case ((string_drop s23310 s23320)) of s23330 => (case ((sep_matches_prefix s23330)) of Some ((_, s23340)) => (case ((string_drop s23330 s23340)) of s23350 => (case ((frm_mnemonic_matches_prefix s23350)) of Some ((rm, s23360)) => (let p00 = (string_drop s23350 s23360) in if (((p00 = ('''')))) then Some (FCVT_S_WU, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s23230 :: " string " \ \\val _s2304_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2304 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2304 s23060 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s23060)) of Some ((FCVT_S_W, s23070)) => (case ((string_drop s23060 s23070)) of s23080 => (case ((spc_matches_prefix0 s23080)) of Some ((_, s23090)) => (case ((string_drop s23080 s23090)) of s23100 => (case ((freg_name_matches_prefix s23100 :: (( 5 Word.word * ii)) option)) of Some ((rd, s23110)) => (case ((string_drop s23100 s23110)) of s23120 => (case ((sep_matches_prefix s23120)) of Some ((_, s23130)) => (case ((string_drop s23120 s23130)) of s23140 => (case ((reg_name_matches_prefix s23140 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s23150)) => (case ((string_drop s23140 s23150)) of s23160 => (case ((sep_matches_prefix s23160)) of Some ((_, s23170)) => (case ((string_drop s23160 s23170)) of s23180 => (case ((frm_mnemonic_matches_prefix s23180)) of Some ((rm, s23190)) => (let p00 = (string_drop s23180 s23190) in if (((p00 = ('''')))) then Some (FCVT_S_W, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s23060 :: " string " \ \\val _s2287_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2287 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2287 s22890 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s22890)) of Some ((FCVT_WU_S, s22900)) => (case ((string_drop s22890 s22900)) of s22910 => (case ((spc_matches_prefix0 s22910)) of Some ((_, s22920)) => (case ((string_drop s22910 s22920)) of s22930 => (case ((reg_name_matches_prefix s22930 :: (( 5 Word.word * ii)) option)) of Some ((rd, s22940)) => (case ((string_drop s22930 s22940)) of s22950 => (case ((sep_matches_prefix s22950)) of Some ((_, s22960)) => (case ((string_drop s22950 s22960)) of s22970 => (case ((freg_name_matches_prefix s22970 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s22980)) => (case ((string_drop s22970 s22980)) of s22990 => (case ((sep_matches_prefix s22990)) of Some ((_, s23000)) => (case ((string_drop s22990 s23000)) of s23010 => (case ((frm_mnemonic_matches_prefix s23010)) of Some ((rm, s23020)) => (let p00 = (string_drop s23010 s23020) in if (((p00 = ('''')))) then Some (FCVT_WU_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s22890 :: " string " \ \\val _s2270_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2270 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2270 s22720 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s22720)) of Some ((FCVT_W_S, s22730)) => (case ((string_drop s22720 s22730)) of s22740 => (case ((spc_matches_prefix0 s22740)) of Some ((_, s22750)) => (case ((string_drop s22740 s22750)) of s22760 => (case ((reg_name_matches_prefix s22760 :: (( 5 Word.word * ii)) option)) of Some ((rd, s22770)) => (case ((string_drop s22760 s22770)) of s22780 => (case ((sep_matches_prefix s22780)) of Some ((_, s22790)) => (case ((string_drop s22780 s22790)) of s22800 => (case ((freg_name_matches_prefix s22800 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s22810)) => (case ((string_drop s22800 s22810)) of s22820 => (case ((sep_matches_prefix s22820)) of Some ((_, s22830)) => (case ((string_drop s22820 s22830)) of s22840 => (case ((frm_mnemonic_matches_prefix s22840)) of Some ((rm, s22850)) => (let p00 = (string_drop s22840 s22850) in if (((p00 = ('''')))) then Some (FCVT_W_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s22720 :: " string " \ \\val _s2253_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s2253 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2253 s22550 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s22550)) of Some ((FSQRT_S, s22560)) => (case ((string_drop s22550 s22560)) of s22570 => (case ((spc_matches_prefix0 s22570)) of Some ((_, s22580)) => (case ((string_drop s22570 s22580)) of s22590 => (case ((freg_name_matches_prefix s22590 :: (( 5 Word.word * ii)) option)) of Some ((rd, s22600)) => (case ((string_drop s22590 s22600)) of s22610 => (case ((sep_matches_prefix s22610)) of Some ((_, s22620)) => (case ((string_drop s22610 s22620)) of s22630 => (case ((freg_name_matches_prefix s22630 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s22640)) => (case ((string_drop s22630 s22640)) of s22650 => (case ((sep_matches_prefix s22650)) of Some ((_, s22660)) => (case ((string_drop s22650 s22660)) of s22670 => (case ((frm_mnemonic_matches_prefix s22670)) of Some ((rm, s22680)) => (let p00 = (string_drop s22670 s22680) in if (((p00 = ('''')))) then Some (FSQRT_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s22550 :: " string " \ \\val _s2232_ : string -> maybe ((f_bin_rm_op_S * mword ty5 * mword ty5 * mword ty5 * rounding_mode))\\ definition s2232 :: \ string \(f_bin_rm_op_S*(5)Word.word*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2232 s22340 = ( (case ((f_bin_rm_type_mnemonic_S_matches_prefix s22340)) of Some ((op1, s22350)) => (case ((string_drop s22340 s22350)) of s22360 => (case ((spc_matches_prefix0 s22360)) of Some ((_, s22370)) => (case ((string_drop s22360 s22370)) of s22380 => (case ((freg_name_matches_prefix s22380 :: (( 5 Word.word * ii)) option)) of Some ((rd, s22390)) => (case ((string_drop s22380 s22390)) of s22400 => (case ((sep_matches_prefix s22400)) of Some ((_, s22410)) => (case ((string_drop s22400 s22410)) of s22420 => (case ((freg_name_matches_prefix s22420 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s22430)) => (case ((string_drop s22420 s22430)) of s22440 => (case ((sep_matches_prefix s22440)) of Some ((_, s22450)) => (case ((string_drop s22440 s22450)) of s22460 => (case ((freg_name_matches_prefix s22460 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s22470)) => (case ((string_drop s22460 s22470)) of s22480 => (case ((sep_matches_prefix s22480)) of Some ((_, s22490)) => (case ((string_drop s22480 s22490)) of s22500 => (case ((frm_mnemonic_matches_prefix s22500)) of Some ((rm, s22510)) => (let p00 = (string_drop s22500 s22510) in if (((p00 = ('''')))) then Some (op1, rd, rs1, rs2, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s22340 :: " string " \ \\val _s2207_ : string -> maybe ((f_madd_op_S * mword ty5 * mword ty5 * mword ty5 * mword ty5 * rounding_mode))\\ definition s2207 :: \ string \(f_madd_op_S*(5)Word.word*(5)Word.word*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s2207 s22090 = ( (case ((f_madd_type_mnemonic_S_matches_prefix s22090)) of Some ((op1, s22100)) => (case ((string_drop s22090 s22100)) of s22110 => (case ((spc_matches_prefix0 s22110)) of Some ((_, s22120)) => (case ((string_drop s22110 s22120)) of s22130 => (case ((freg_name_matches_prefix s22130 :: (( 5 Word.word * ii)) option)) of Some ((rd, s22140)) => (case ((string_drop s22130 s22140)) of s22150 => (case ((sep_matches_prefix s22150)) of Some ((_, s22160)) => (case ((string_drop s22150 s22160)) of s22170 => (case ((freg_name_matches_prefix s22170 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s22180)) => (case ((string_drop s22170 s22180)) of s22190 => (case ((sep_matches_prefix s22190)) of Some ((_, s22200)) => (case ((string_drop s22190 s22200)) of s22210 => (case ((freg_name_matches_prefix s22210 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s22220)) => (case ((string_drop s22210 s22220)) of s22230 => (case ((sep_matches_prefix s22230)) of Some ((_, s22240)) => (case ((string_drop s22230 s22240)) of s22250 => (case ((freg_name_matches_prefix s22250 :: (( 5 Word.word * ii)) option)) of Some ((rs3, s22260)) => (case ((string_drop s22250 s22260)) of s22270 => (case ((sep_matches_prefix s22270)) of Some ((_, s22280)) => (case ((string_drop s22270 s22280)) of s22290 => (case ((frm_mnemonic_matches_prefix s22290)) of Some ((rm, s22300)) => (let p00 = (string_drop s22290 s22300) in if (((p00 = ('''')))) then Some (op1, rd, rs1, rs2, rs3, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s22090 :: " string " \ \\val _s2183_ : string -> maybe ((word_width * mword ty5 * mword ty12 * mword ty5))\\ definition s2183 :: \ string \(word_width*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s2183 s21840 = ( (let s21850 = s21840 in if ((string_startswith s21850 (''fs''))) then (case ((string_drop s21850 ((string_length (''fs''))))) of s21860 => (case ((size_mnemonic_matches_prefix s21860)) of Some ((width, s21870)) => (case ((string_drop s21860 s21870)) of s21880 => (case ((spc_matches_prefix0 s21880)) of Some ((_, s21890)) => (case ((string_drop s21880 s21890)) of s21900 => (case ((freg_name_matches_prefix s21900 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s21910)) => (case ((string_drop s21900 s21910)) of s21920 => (case ((sep_matches_prefix s21920)) of Some ((_, s21930)) => (case ((string_drop s21920 s21930)) of s21940 => (case ((hex_bits_12_matches_prefix0 s21940 :: (( 12 Word.word * ii)) option)) of Some ((imm, s21950)) => (case ((string_drop s21940 s21950)) of s21960 => (case ((opt_spc_matches_prefix0 s21960)) of Some ((_, s21970)) => (let s21980 = (string_drop s21960 s21970) in if ((string_startswith s21980 (''(''))) then (case ((string_drop s21980 ((string_length (''(''))))) of s21990 => (case ((opt_spc_matches_prefix0 s21990)) of Some ((_, s22000)) => (case ((string_drop s21990 s22000)) of s22010 => (case ((reg_name_matches_prefix s22010 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s22020)) => (case ((string_drop s22010 s22020)) of s22030 => (case ((opt_spc_matches_prefix0 s22030)) of Some ((_, s22040)) => (let s22050 = (string_drop s22030 s22040) in if ((string_startswith s22050 ('')''))) then (let p00 = (string_drop s22050 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (width, rs2, imm, rs1) else None) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s21840 :: " string " \ \\val _s2159_ : string -> maybe ((word_width * mword ty5 * mword ty12 * mword ty5))\\ definition s2159 :: \ string \(word_width*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s2159 s21600 = ( (let s21610 = s21600 in if ((string_startswith s21610 (''fl''))) then (case ((string_drop s21610 ((string_length (''fl''))))) of s21620 => (case ((size_mnemonic_matches_prefix s21620)) of Some ((width, s21630)) => (case ((string_drop s21620 s21630)) of s21640 => (case ((spc_matches_prefix0 s21640)) of Some ((_, s21650)) => (case ((string_drop s21640 s21650)) of s21660 => (case ((freg_name_matches_prefix s21660 :: (( 5 Word.word * ii)) option)) of Some ((rd, s21670)) => (case ((string_drop s21660 s21670)) of s21680 => (case ((sep_matches_prefix s21680)) of Some ((_, s21690)) => (case ((string_drop s21680 s21690)) of s21700 => (case ((hex_bits_12_matches_prefix0 s21700 :: (( 12 Word.word * ii)) option)) of Some ((imm, s21710)) => (case ((string_drop s21700 s21710)) of s21720 => (case ((opt_spc_matches_prefix0 s21720)) of Some ((_, s21730)) => (let s21740 = (string_drop s21720 s21730) in if ((string_startswith s21740 (''(''))) then (case ((string_drop s21740 ((string_length (''(''))))) of s21750 => (case ((opt_spc_matches_prefix0 s21750)) of Some ((_, s21760)) => (case ((string_drop s21750 s21760)) of s21770 => (case ((reg_name_matches_prefix s21770 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s21780)) => (case ((string_drop s21770 s21780)) of s21790 => (case ((opt_spc_matches_prefix0 s21790)) of Some ((_, s21800)) => (let s21810 = (string_drop s21790 s21800) in if ((string_startswith s21810 ('')''))) then (let p00 = (string_drop s21810 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (width, rd, imm, rs1) else None) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s21600 :: " string " \ \\val _s2147_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12))\\ definition s2147 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word)option \ where \ s2147 s21480 = ( (let s21490 = s21480 in if ((string_startswith s21490 (''fence.i.reserved.''))) then (case ((string_drop s21490 ((string_length (''fence.i.reserved.''))))) of s21500 => (case ((reg_name_matches_prefix s21500 :: (( 5 Word.word * ii)) option)) of Some ((rd, s21510)) => (let s21520 = (string_drop s21500 s21510) in if ((string_startswith s21520 (''.''))) then (case ((string_drop s21520 ((string_length (''.''))))) of s21530 => (case ((reg_name_matches_prefix s21530 :: (( 5 Word.word * ii)) option)) of Some ((rs, s21540)) => (let s21550 = (string_drop s21530 s21540) in if ((string_startswith s21550 (''.''))) then (case ((string_drop s21550 ((string_length (''.''))))) of s21560 => (case ((hex_bits_12_matches_prefix0 s21560 :: (( 12 Word.word * ii)) option)) of Some ((imm, s21570)) => (let p00 = (string_drop s21560 s21570) in if (((p00 = ('''')))) then Some (rd, rs, imm) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None))\ for s21480 :: " string " \ \\val _s2129_ : string -> maybe ((mword ty4 * mword ty4 * mword ty5 * mword ty5 * mword ty4))\\ definition s2129 :: \ string \((4)Word.word*(4)Word.word*(5)Word.word*(5)Word.word*(4)Word.word)option \ where \ s2129 s21300 = ( (let s21310 = s21300 in if ((string_startswith s21310 (''fence.reserved.''))) then (case ((string_drop s21310 ((string_length (''fence.reserved.''))))) of s21320 => (case ((fence_bits_matches_prefix s21320 :: (( 4 Word.word * ii)) option)) of Some ((pred, s21330)) => (let s21340 = (string_drop s21320 s21330) in if ((string_startswith s21340 (''.''))) then (case ((string_drop s21340 ((string_length (''.''))))) of s21350 => (case ((fence_bits_matches_prefix s21350 :: (( 4 Word.word * ii)) option)) of Some ((succ, s21360)) => (let s21370 = (string_drop s21350 s21360) in if ((string_startswith s21370 (''.''))) then (case ((string_drop s21370 ((string_length (''.''))))) of s21380 => (case ((reg_name_matches_prefix s21380 :: (( 5 Word.word * ii)) option)) of Some ((rs, s21390)) => (let s21400 = (string_drop s21380 s21390) in if ((string_startswith s21400 (''.''))) then (case ((string_drop s21400 ((string_length (''.''))))) of s21410 => (case ((reg_name_matches_prefix s21410 :: (( 5 Word.word * ii)) option)) of Some ((rd, s21420)) => (let s21430 = (string_drop s21410 s21420) in if ((string_startswith s21430 (''.''))) then (case ((string_drop s21430 ((string_length (''.''))))) of s21440 => (case ((hex_bits_4_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s21440 :: (( 4 Word.word * ii)) option)) of Some ((fm, s21450)) => (let p00 = (string_drop s21440 s21450) in if (((p00 = ('''')))) then Some (pred, succ, rs, rd, fm) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None))\ for s21300 :: " string " \ \\val _s2123_ : string -> maybe (mword ty3)\\ definition s2123 :: \ string \((3)Word.word)option \ where \ s2123 s21240 = ( (let s21250 = s21240 in if ((string_startswith s21250 (''c.srai.hint.''))) then (case ((string_drop s21250 ((string_length (''c.srai.hint.''))))) of s21260 => (case ((creg_name_matches_prefix s21260 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s21270)) => (let p00 = (string_drop s21260 s21270) in if (((p00 = ('''')))) then Some rsd else None) | _ => None ) ) else None))\ for s21240 :: " string " \ \\val _s2117_ : string -> maybe (mword ty3)\\ definition s2117 :: \ string \((3)Word.word)option \ where \ s2117 s21180 = ( (let s21190 = s21180 in if ((string_startswith s21190 (''c.srli.hint.''))) then (case ((string_drop s21190 ((string_length (''c.srli.hint.''))))) of s21200 => (case ((creg_name_matches_prefix s21200 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s21210)) => (let p00 = (string_drop s21200 s21210) in if (((p00 = ('''')))) then Some rsd else None) | _ => None ) ) else None))\ for s21180 :: " string " \ \\val _s2108_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s2108 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s2108 s21090 = ( (let s21100 = s21090 in if ((string_startswith s21100 (''c.slli.hint.''))) then (case ((string_drop s21100 ((string_length (''c.slli.hint.''))))) of s21110 => (case ((reg_name_matches_prefix s21110 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s21120)) => (let s21130 = (string_drop s21110 s21120) in if ((string_startswith s21130 (''.''))) then (case ((string_drop s21130 ((string_length (''.''))))) of s21140 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s21140 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s21150)) => (let p00 = (string_drop s21140 s21150) in if (((p00 = ('''')))) then Some (rsd, shamt) else None) | _ => None ) ) else None) | _ => None ) ) else None))\ for s21090 :: " string " \ \\val _s2102_ : string -> maybe (mword ty5)\\ definition s2102 :: \ string \((5)Word.word)option \ where \ s2102 s21030 = ( (let s21040 = s21030 in if ((string_startswith s21040 (''c.add.hint.''))) then (case ((string_drop s21040 ((string_length (''c.add.hint.''))))) of s21050 => (case ((reg_name_matches_prefix s21050 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s21060)) => (let p00 = (string_drop s21050 s21060) in if (((p00 = ('''')))) then Some rs2 else None) | _ => None ) ) else None))\ for s21030 :: " string " \ \\val _s2096_ : string -> maybe (mword ty5)\\ definition s2096 :: \ string \((5)Word.word)option \ where \ s2096 s20970 = ( (let s20980 = s20970 in if ((string_startswith s20980 (''c.mv.hint.''))) then (case ((string_drop s20980 ((string_length (''c.mv.hint.''))))) of s20990 => (case ((reg_name_matches_prefix s20990 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s21000)) => (let p00 = (string_drop s20990 s21000) in if (((p00 = ('''')))) then Some rs2 else None) | _ => None ) ) else None))\ for s20970 :: " string " \ \\val _s2090_ : string -> maybe (mword ty6)\\ definition s2090 :: \ string \((6)Word.word)option \ where \ s2090 s20910 = ( (let s20920 = s20910 in if ((string_startswith s20920 (''c.lui.hint.''))) then (case ((string_drop s20920 ((string_length (''c.lui.hint.''))))) of s20930 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s20930 :: (( 6 Word.word * ii)) option)) of Some ((imm, s20940)) => (let p00 = (string_drop s20930 s20940) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) else None))\ for s20910 :: " string " \ \\val _s2084_ : string -> maybe (mword ty6)\\ definition s2084 :: \ string \((6)Word.word)option \ where \ s2084 s20850 = ( (let s20860 = s20850 in if ((string_startswith s20860 (''c.li.hint.''))) then (case ((string_drop s20860 ((string_length (''c.li.hint.''))))) of s20870 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s20870 :: (( 6 Word.word * ii)) option)) of Some ((imm, s20880)) => (let p00 = (string_drop s20870 s20880) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) else None))\ for s20850 :: " string " \ \\val _s2078_ : string -> maybe (mword ty5)\\ definition s2078 :: \ string \((5)Word.word)option \ where \ s2078 s20790 = ( (let s20800 = s20790 in if ((string_startswith s20800 (''c.addi.hint.''))) then (case ((string_drop s20800 ((string_length (''c.addi.hint.''))))) of s20810 => (case ((reg_name_matches_prefix s20810 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s20820)) => (let p00 = (string_drop s20810 s20820) in if (((p00 = ('''')))) then Some rsd else None) | _ => None ) ) else None))\ for s20790 :: " string " \ \\val _s2072_ : string -> maybe (mword ty6)\\ definition s2072 :: \ string \((6)Word.word)option \ where \ s2072 s20730 = ( (let s20740 = s20730 in if ((string_startswith s20740 (''c.nop.hint.''))) then (case ((string_drop s20740 ((string_length (''c.nop.hint.''))))) of s20750 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s20750 :: (( 6 Word.word * ii)) option)) of Some ((imm, s20760)) => (let p00 = (string_drop s20750 s20760) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) else None))\ for s20730 :: " string " \ \\val _s2055_ : string -> maybe ((csrop * mword ty5 * mword ty12 * mword ty5))\\ definition s2055 :: \ string \(csrop*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s2055 s20570 = ( (case ((csr_mnemonic_matches_prefix s20570)) of Some ((op1, s20580)) => (case ((string_drop s20570 s20580)) of s20590 => (case ((spc_matches_prefix0 s20590)) of Some ((_, s20600)) => (case ((string_drop s20590 s20600)) of s20610 => (case ((reg_name_matches_prefix s20610 :: (( 5 Word.word * ii)) option)) of Some ((rd, s20620)) => (case ((string_drop s20610 s20620)) of s20630 => (case ((sep_matches_prefix s20630)) of Some ((_, s20640)) => (case ((string_drop s20630 s20640)) of s20650 => (case ((csr_name_map_matches_prefix s20650 :: (( 12 Word.word * ii)) option)) of Some ((csr, s20660)) => (case ((string_drop s20650 s20660)) of s20670 => (case ((sep_matches_prefix s20670)) of Some ((_, s20680)) => (case ((string_drop s20670 s20680)) of s20690 => (case ((reg_name_matches_prefix s20690 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s20700)) => (let p00 = (string_drop s20690 s20700) in if (((p00 = ('''')))) then Some (op1, rd, csr, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s20570 :: " string " \ \\val _s2037_ : string -> maybe ((csrop * mword ty5 * mword ty12 * mword ty5))\\ definition s2037 :: \ string \(csrop*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s2037 s20390 = ( (case ((csr_mnemonic_matches_prefix s20390)) of Some ((op1, s20400)) => (let s20410 = (string_drop s20390 s20400) in if ((string_startswith s20410 (''i''))) then (case ((string_drop s20410 ((string_length (''i''))))) of s20420 => (case ((spc_matches_prefix0 s20420)) of Some ((_, s20430)) => (case ((string_drop s20420 s20430)) of s20440 => (case ((reg_name_matches_prefix s20440 :: (( 5 Word.word * ii)) option)) of Some ((rd, s20450)) => (case ((string_drop s20440 s20450)) of s20460 => (case ((sep_matches_prefix s20460)) of Some ((_, s20470)) => (case ((string_drop s20460 s20470)) of s20480 => (case ((csr_name_map_matches_prefix s20480 :: (( 12 Word.word * ii)) option)) of Some ((csr, s20490)) => (case ((string_drop s20480 s20490)) of s20500 => (case ((sep_matches_prefix s20500)) of Some ((_, s20510)) => (case ((string_drop s20500 s20510)) of s20520 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s20520 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s20530)) => (let p00 = (string_drop s20520 s20530) in if (((p00 = ('''')))) then Some (op1, rd, csr, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ))\ for s20390 :: " string " \ \\val _s2018_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5))\\ definition s2018 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2018 s20190 = ( (let s20200 = s20190 in if ((string_startswith s20200 (''rem''))) then (case ((string_drop s20200 ((string_length (''rem''))))) of s20210 => (case ((maybe_not_u_matches_prefix s20210)) of Some ((s, s20220)) => (let s20230 = (string_drop s20210 s20220) in if ((string_startswith s20230 (''w''))) then (case ((string_drop s20230 ((string_length (''w''))))) of s20240 => (case ((spc_matches_prefix0 s20240)) of Some ((_, s20250)) => (case ((string_drop s20240 s20250)) of s20260 => (case ((reg_name_matches_prefix s20260 :: (( 5 Word.word * ii)) option)) of Some ((rd, s20270)) => (case ((string_drop s20260 s20270)) of s20280 => (case ((sep_matches_prefix s20280)) of Some ((_, s20290)) => (case ((string_drop s20280 s20290)) of s20300 => (case ((reg_name_matches_prefix s20300 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s20310)) => (case ((string_drop s20300 s20310)) of s20320 => (case ((sep_matches_prefix s20320)) of Some ((_, s20330)) => (case ((string_drop s20320 s20330)) of s20340 => (case ((reg_name_matches_prefix s20340 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s20350)) => (let p00 = (string_drop s20340 s20350) in if (((p00 = ('''')))) then Some (s, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) else None))\ for s20190 :: " string " \ \\val _s1999_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5))\\ definition s1999 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1999 s20000 = ( (let s20010 = s20000 in if ((string_startswith s20010 (''div''))) then (case ((string_drop s20010 ((string_length (''div''))))) of s20020 => (case ((maybe_not_u_matches_prefix s20020)) of Some ((s, s20030)) => (let s20040 = (string_drop s20020 s20030) in if ((string_startswith s20040 (''w''))) then (case ((string_drop s20040 ((string_length (''w''))))) of s20050 => (case ((spc_matches_prefix0 s20050)) of Some ((_, s20060)) => (case ((string_drop s20050 s20060)) of s20070 => (case ((reg_name_matches_prefix s20070 :: (( 5 Word.word * ii)) option)) of Some ((rd, s20080)) => (case ((string_drop s20070 s20080)) of s20090 => (case ((sep_matches_prefix s20090)) of Some ((_, s20100)) => (case ((string_drop s20090 s20100)) of s20110 => (case ((reg_name_matches_prefix s20110 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s20120)) => (case ((string_drop s20110 s20120)) of s20130 => (case ((sep_matches_prefix s20130)) of Some ((_, s20140)) => (case ((string_drop s20130 s20140)) of s20150 => (case ((reg_name_matches_prefix s20150 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s20160)) => (let p00 = (string_drop s20150 s20160) in if (((p00 = ('''')))) then Some (s, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) else None))\ for s20000 :: " string " \ \\val _s1983_ : string -> maybe ((mword ty5 * mword ty5 * mword ty5))\\ definition s1983 :: \ string \((5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1983 s19840 = ( (let s19850 = s19840 in if ((string_startswith s19850 (''mulw''))) then (case ((string_drop s19850 ((string_length (''mulw''))))) of s19860 => (case ((spc_matches_prefix0 s19860)) of Some ((_, s19870)) => (case ((string_drop s19860 s19870)) of s19880 => (case ((reg_name_matches_prefix s19880 :: (( 5 Word.word * ii)) option)) of Some ((rd, s19890)) => (case ((string_drop s19880 s19890)) of s19900 => (case ((sep_matches_prefix s19900)) of Some ((_, s19910)) => (case ((string_drop s19900 s19910)) of s19920 => (case ((reg_name_matches_prefix s19920 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s19930)) => (case ((string_drop s19920 s19930)) of s19940 => (case ((sep_matches_prefix s19940)) of Some ((_, s19950)) => (case ((string_drop s19940 s19950)) of s19960 => (case ((reg_name_matches_prefix s19960 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s19970)) => (let p00 = (string_drop s19960 s19970) in if (((p00 = ('''')))) then Some (rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s19840 :: " string " \ \\val _s1965_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5))\\ definition s1965 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1965 s19660 = ( (let s19670 = s19660 in if ((string_startswith s19670 (''rem''))) then (case ((string_drop s19670 ((string_length (''rem''))))) of s19680 => (case ((maybe_not_u_matches_prefix s19680)) of Some ((s, s19690)) => (case ((string_drop s19680 s19690)) of s19700 => (case ((spc_matches_prefix0 s19700)) of Some ((_, s19710)) => (case ((string_drop s19700 s19710)) of s19720 => (case ((reg_name_matches_prefix s19720 :: (( 5 Word.word * ii)) option)) of Some ((rd, s19730)) => (case ((string_drop s19720 s19730)) of s19740 => (case ((sep_matches_prefix s19740)) of Some ((_, s19750)) => (case ((string_drop s19740 s19750)) of s19760 => (case ((reg_name_matches_prefix s19760 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s19770)) => (case ((string_drop s19760 s19770)) of s19780 => (case ((sep_matches_prefix s19780)) of Some ((_, s19790)) => (case ((string_drop s19780 s19790)) of s19800 => (case ((reg_name_matches_prefix s19800 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s19810)) => (let p00 = (string_drop s19800 s19810) in if (((p00 = ('''')))) then Some (s, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s19660 :: " string " \ \\val _s1947_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5))\\ definition s1947 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1947 s19480 = ( (let s19490 = s19480 in if ((string_startswith s19490 (''div''))) then (case ((string_drop s19490 ((string_length (''div''))))) of s19500 => (case ((maybe_not_u_matches_prefix s19500)) of Some ((s, s19510)) => (case ((string_drop s19500 s19510)) of s19520 => (case ((spc_matches_prefix0 s19520)) of Some ((_, s19530)) => (case ((string_drop s19520 s19530)) of s19540 => (case ((reg_name_matches_prefix s19540 :: (( 5 Word.word * ii)) option)) of Some ((rd, s19550)) => (case ((string_drop s19540 s19550)) of s19560 => (case ((sep_matches_prefix s19560)) of Some ((_, s19570)) => (case ((string_drop s19560 s19570)) of s19580 => (case ((reg_name_matches_prefix s19580 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s19590)) => (case ((string_drop s19580 s19590)) of s19600 => (case ((sep_matches_prefix s19600)) of Some ((_, s19610)) => (case ((string_drop s19600 s19610)) of s19620 => (case ((reg_name_matches_prefix s19620 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s19630)) => (let p00 = (string_drop s19620 s19630) in if (((p00 = ('''')))) then Some (s, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s19480 :: " string " \ \\val _s1930_ : string -> maybe ((bool * bool * bool * mword ty5 * mword ty5 * mword ty5))\\ definition s1930 :: \ string \(bool*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1930 s19320 = ( (case ((mul_mnemonic_matches_prefix s19320)) of Some (((high, signed1, signed2), s19330)) => (case ((string_drop s19320 s19330)) of s19340 => (case ((spc_matches_prefix0 s19340)) of Some ((_, s19350)) => (case ((string_drop s19340 s19350)) of s19360 => (case ((reg_name_matches_prefix s19360 :: (( 5 Word.word * ii)) option)) of Some ((rd, s19370)) => (case ((string_drop s19360 s19370)) of s19380 => (case ((sep_matches_prefix s19380)) of Some ((_, s19390)) => (case ((string_drop s19380 s19390)) of s19400 => (case ((reg_name_matches_prefix s19400 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s19410)) => (case ((string_drop s19400 s19410)) of s19420 => (case ((sep_matches_prefix s19420)) of Some ((_, s19430)) => (case ((string_drop s19420 s19430)) of s19440 => (case ((reg_name_matches_prefix s19440 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s19450)) => (let p00 = (string_drop s19440 s19450) in if (((p00 = ('''')))) then Some (high, signed1, signed2, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s19320 :: " string " \ \\val _s1918_ : string -> maybe ((mword ty5 * mword ty5))\\ definition s1918 :: \ string \((5)Word.word*(5)Word.word)option \ where \ s1918 s19190 = ( (let s19200 = s19190 in if ((string_startswith s19200 (''c.add''))) then (case ((string_drop s19200 ((string_length (''c.add''))))) of s19210 => (case ((spc_matches_prefix0 s19210)) of Some ((_, s19220)) => (case ((string_drop s19210 s19220)) of s19230 => (case ((reg_name_matches_prefix s19230 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s19240)) => (case ((string_drop s19230 s19240)) of s19250 => (case ((sep_matches_prefix s19250)) of Some ((_, s19260)) => (case ((string_drop s19250 s19260)) of s19270 => (case ((reg_name_matches_prefix s19270 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s19280)) => (let p00 = (string_drop s19270 s19280) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s19190 :: " string " \ \\val _s1906_ : string -> maybe ((mword ty5 * mword ty5))\\ definition s1906 :: \ string \((5)Word.word*(5)Word.word)option \ where \ s1906 s19070 = ( (let s19080 = s19070 in if ((string_startswith s19080 (''c.mv''))) then (case ((string_drop s19080 ((string_length (''c.mv''))))) of s19090 => (case ((spc_matches_prefix0 s19090)) of Some ((_, s19100)) => (case ((string_drop s19090 s19100)) of s19110 => (case ((reg_name_matches_prefix s19110 :: (( 5 Word.word * ii)) option)) of Some ((rd, s19120)) => (case ((string_drop s19110 s19120)) of s19130 => (case ((sep_matches_prefix s19130)) of Some ((_, s19140)) => (case ((string_drop s19130 s19140)) of s19150 => (case ((reg_name_matches_prefix s19150 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s19160)) => (let p00 = (string_drop s19150 s19160) in if (((p00 = ('''')))) then Some (rd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s19070 :: " string " \ \\val _s1898_ : string -> maybe (mword ty5)\\ definition s1898 :: \ string \((5)Word.word)option \ where \ s1898 s18990 = ( (let s19000 = s18990 in if ((string_startswith s19000 (''c.jalr''))) then (case ((string_drop s19000 ((string_length (''c.jalr''))))) of s19010 => (case ((spc_matches_prefix0 s19010)) of Some ((_, s19020)) => (case ((string_drop s19010 s19020)) of s19030 => (case ((reg_name_matches_prefix s19030 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s19040)) => (let p00 = (string_drop s19030 s19040) in if (((p00 = ('''')))) then Some rs1 else None) | _ => None ) ) | _ => None ) ) else None))\ for s18990 :: " string " \ \\val _s1890_ : string -> maybe (mword ty5)\\ definition s1890 :: \ string \((5)Word.word)option \ where \ s1890 s18910 = ( (let s18920 = s18910 in if ((string_startswith s18920 (''c.jr''))) then (case ((string_drop s18920 ((string_length (''c.jr''))))) of s18930 => (case ((spc_matches_prefix0 s18930)) of Some ((_, s18940)) => (case ((string_drop s18930 s18940)) of s18950 => (case ((reg_name_matches_prefix s18950 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s18960)) => (let p00 = (string_drop s18950 s18960) in if (((p00 = ('''')))) then Some rs1 else None) | _ => None ) ) | _ => None ) ) else None))\ for s18910 :: " string " \ \\val _s1878_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1878 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1878 s18790 = ( (let s18800 = s18790 in if ((string_startswith s18800 (''c.sdsp''))) then (case ((string_drop s18800 ((string_length (''c.sdsp''))))) of s18810 => (case ((spc_matches_prefix0 s18810)) of Some ((_, s18820)) => (case ((string_drop s18810 s18820)) of s18830 => (case ((reg_name_matches_prefix s18830 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s18840)) => (case ((string_drop s18830 s18840)) of s18850 => (case ((sep_matches_prefix s18850)) of Some ((_, s18860)) => (case ((string_drop s18850 s18860)) of s18870 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s18870 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s18880)) => (let p00 = (string_drop s18870 s18880) in if (((p00 = ('''')))) then Some (rs2, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s18790 :: " string " \ \\val _s1866_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1866 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1866 s18670 = ( (let s18680 = s18670 in if ((string_startswith s18680 (''c.swsp''))) then (case ((string_drop s18680 ((string_length (''c.swsp''))))) of s18690 => (case ((spc_matches_prefix0 s18690)) of Some ((_, s18700)) => (case ((string_drop s18690 s18700)) of s18710 => (case ((reg_name_matches_prefix s18710 :: (( 5 Word.word * ii)) option)) of Some ((rd, s18720)) => (case ((string_drop s18710 s18720)) of s18730 => (case ((sep_matches_prefix s18730)) of Some ((_, s18740)) => (case ((string_drop s18730 s18740)) of s18750 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s18750 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s18760)) => (let p00 = (string_drop s18750 s18760) in if (((p00 = ('''')))) then Some (rd, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s18670 :: " string " \ \\val _s1854_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1854 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1854 s18550 = ( (let s18560 = s18550 in if ((string_startswith s18560 (''c.ldsp''))) then (case ((string_drop s18560 ((string_length (''c.ldsp''))))) of s18570 => (case ((spc_matches_prefix0 s18570)) of Some ((_, s18580)) => (case ((string_drop s18570 s18580)) of s18590 => (case ((reg_name_matches_prefix s18590 :: (( 5 Word.word * ii)) option)) of Some ((rd, s18600)) => (case ((string_drop s18590 s18600)) of s18610 => (case ((sep_matches_prefix s18610)) of Some ((_, s18620)) => (case ((string_drop s18610 s18620)) of s18630 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s18630 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s18640)) => (let p00 = (string_drop s18630 s18640) in if (((p00 = ('''')))) then Some (rd, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s18550 :: " string " \ \\val _s1842_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1842 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1842 s18430 = ( (let s18440 = s18430 in if ((string_startswith s18440 (''c.lwsp''))) then (case ((string_drop s18440 ((string_length (''c.lwsp''))))) of s18450 => (case ((spc_matches_prefix0 s18450)) of Some ((_, s18460)) => (case ((string_drop s18450 s18460)) of s18470 => (case ((reg_name_matches_prefix s18470 :: (( 5 Word.word * ii)) option)) of Some ((rd, s18480)) => (case ((string_drop s18470 s18480)) of s18490 => (case ((sep_matches_prefix s18490)) of Some ((_, s18500)) => (case ((string_drop s18490 s18500)) of s18510 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s18510 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s18520)) => (let p00 = (string_drop s18510 s18520) in if (((p00 = ('''')))) then Some (rd, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s18430 :: " string " \ \\val _s1830_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1830 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1830 s18310 = ( (let s18320 = s18310 in if ((string_startswith s18320 (''c.slli''))) then (case ((string_drop s18320 ((string_length (''c.slli''))))) of s18330 => (case ((spc_matches_prefix0 s18330)) of Some ((_, s18340)) => (case ((string_drop s18330 s18340)) of s18350 => (case ((reg_name_matches_prefix s18350 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s18360)) => (case ((string_drop s18350 s18360)) of s18370 => (case ((sep_matches_prefix s18370)) of Some ((_, s18380)) => (case ((string_drop s18370 s18380)) of s18390 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s18390 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s18400)) => (let p00 = (string_drop s18390 s18400) in if (((p00 = ('''')))) then Some (rsd, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s18310 :: " string " \ \\val _s1818_ : string -> maybe ((mword ty3 * mword ty8))\\ definition s1818 :: \ string \((3)Word.word*(8)Word.word)option \ where \ s1818 s18190 = ( (let s18200 = s18190 in if ((string_startswith s18200 (''c.bnez''))) then (case ((string_drop s18200 ((string_length (''c.bnez''))))) of s18210 => (case ((spc_matches_prefix0 s18210)) of Some ((_, s18220)) => (case ((string_drop s18210 s18220)) of s18230 => (case ((creg_name_matches_prefix s18230 :: (( 3 Word.word * ii)) option)) of Some ((rs, s18240)) => (case ((string_drop s18230 s18240)) of s18250 => (case ((sep_matches_prefix s18250)) of Some ((_, s18260)) => (case ((string_drop s18250 s18260)) of s18270 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s18270 :: (( 8 Word.word * ii)) option)) of Some ((imm, s18280)) => (let p00 = (string_drop s18270 s18280) in if (((p00 = ('''')))) then Some (rs, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s18190 :: " string " \ \\val _s1806_ : string -> maybe ((mword ty3 * mword ty8))\\ definition s1806 :: \ string \((3)Word.word*(8)Word.word)option \ where \ s1806 s18070 = ( (let s18080 = s18070 in if ((string_startswith s18080 (''c.beqz''))) then (case ((string_drop s18080 ((string_length (''c.beqz''))))) of s18090 => (case ((spc_matches_prefix0 s18090)) of Some ((_, s18100)) => (case ((string_drop s18090 s18100)) of s18110 => (case ((creg_name_matches_prefix s18110 :: (( 3 Word.word * ii)) option)) of Some ((rs, s18120)) => (case ((string_drop s18110 s18120)) of s18130 => (case ((sep_matches_prefix s18130)) of Some ((_, s18140)) => (case ((string_drop s18130 s18140)) of s18150 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s18150 :: (( 8 Word.word * ii)) option)) of Some ((imm, s18160)) => (let p00 = (string_drop s18150 s18160) in if (((p00 = ('''')))) then Some (rs, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s18070 :: " string " \ \\val _s1798_ : string -> maybe (mword ty11)\\ definition s1798 :: \ string \((11)Word.word)option \ where \ s1798 s17990 = ( (let s18000 = s17990 in if ((string_startswith s18000 (''c.j''))) then (case ((string_drop s18000 ((string_length (''c.j''))))) of s18010 => (case ((spc_matches_prefix0 s18010)) of Some ((_, s18020)) => (case ((string_drop s18010 s18020)) of s18030 => (case ((hex_bits_11_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s18030 :: (( 11 Word.word * ii)) option)) of Some ((imm, s18040)) => (let p00 = (string_drop s18030 s18040) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) | _ => None ) ) else None))\ for s17990 :: " string " \ \\val _s1786_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s1786 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s1786 s17870 = ( (let s17880 = s17870 in if ((string_startswith s17880 (''c.addw''))) then (case ((string_drop s17880 ((string_length (''c.addw''))))) of s17890 => (case ((spc_matches_prefix0 s17890)) of Some ((_, s17900)) => (case ((string_drop s17890 s17900)) of s17910 => (case ((creg_name_matches_prefix s17910 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s17920)) => (case ((string_drop s17910 s17920)) of s17930 => (case ((sep_matches_prefix s17930)) of Some ((_, s17940)) => (case ((string_drop s17930 s17940)) of s17950 => (case ((creg_name_matches_prefix s17950 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s17960)) => (let p00 = (string_drop s17950 s17960) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s17870 :: " string " \ \\val _s1774_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s1774 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s1774 s17750 = ( (let s17760 = s17750 in if ((string_startswith s17760 (''c.subw''))) then (case ((string_drop s17760 ((string_length (''c.subw''))))) of s17770 => (case ((spc_matches_prefix0 s17770)) of Some ((_, s17780)) => (case ((string_drop s17770 s17780)) of s17790 => (case ((creg_name_matches_prefix s17790 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s17800)) => (case ((string_drop s17790 s17800)) of s17810 => (case ((sep_matches_prefix s17810)) of Some ((_, s17820)) => (case ((string_drop s17810 s17820)) of s17830 => (case ((creg_name_matches_prefix s17830 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s17840)) => (let p00 = (string_drop s17830 s17840) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s17750 :: " string " \ \\val _s1762_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s1762 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s1762 s17630 = ( (let s17640 = s17630 in if ((string_startswith s17640 (''c.and''))) then (case ((string_drop s17640 ((string_length (''c.and''))))) of s17650 => (case ((spc_matches_prefix0 s17650)) of Some ((_, s17660)) => (case ((string_drop s17650 s17660)) of s17670 => (case ((creg_name_matches_prefix s17670 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s17680)) => (case ((string_drop s17670 s17680)) of s17690 => (case ((sep_matches_prefix s17690)) of Some ((_, s17700)) => (case ((string_drop s17690 s17700)) of s17710 => (case ((creg_name_matches_prefix s17710 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s17720)) => (let p00 = (string_drop s17710 s17720) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s17630 :: " string " \ \\val _s1750_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s1750 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s1750 s17510 = ( (let s17520 = s17510 in if ((string_startswith s17520 (''c.or''))) then (case ((string_drop s17520 ((string_length (''c.or''))))) of s17530 => (case ((spc_matches_prefix0 s17530)) of Some ((_, s17540)) => (case ((string_drop s17530 s17540)) of s17550 => (case ((creg_name_matches_prefix s17550 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s17560)) => (case ((string_drop s17550 s17560)) of s17570 => (case ((sep_matches_prefix s17570)) of Some ((_, s17580)) => (case ((string_drop s17570 s17580)) of s17590 => (case ((creg_name_matches_prefix s17590 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s17600)) => (let p00 = (string_drop s17590 s17600) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s17510 :: " string " \ \\val _s1738_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s1738 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s1738 s17390 = ( (let s17400 = s17390 in if ((string_startswith s17400 (''c.xor''))) then (case ((string_drop s17400 ((string_length (''c.xor''))))) of s17410 => (case ((spc_matches_prefix0 s17410)) of Some ((_, s17420)) => (case ((string_drop s17410 s17420)) of s17430 => (case ((creg_name_matches_prefix s17430 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s17440)) => (case ((string_drop s17430 s17440)) of s17450 => (case ((sep_matches_prefix s17450)) of Some ((_, s17460)) => (case ((string_drop s17450 s17460)) of s17470 => (case ((creg_name_matches_prefix s17470 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s17480)) => (let p00 = (string_drop s17470 s17480) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s17390 :: " string " \ \\val _s1726_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s1726 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s1726 s17270 = ( (let s17280 = s17270 in if ((string_startswith s17280 (''c.sub''))) then (case ((string_drop s17280 ((string_length (''c.sub''))))) of s17290 => (case ((spc_matches_prefix0 s17290)) of Some ((_, s17300)) => (case ((string_drop s17290 s17300)) of s17310 => (case ((creg_name_matches_prefix s17310 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s17320)) => (case ((string_drop s17310 s17320)) of s17330 => (case ((sep_matches_prefix s17330)) of Some ((_, s17340)) => (case ((string_drop s17330 s17340)) of s17350 => (case ((creg_name_matches_prefix s17350 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s17360)) => (let p00 = (string_drop s17350 s17360) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s17270 :: " string " \ \\val _s1714_ : string -> maybe ((mword ty3 * mword ty6))\\ definition s1714 :: \ string \((3)Word.word*(6)Word.word)option \ where \ s1714 s17150 = ( (let s17160 = s17150 in if ((string_startswith s17160 (''c.andi''))) then (case ((string_drop s17160 ((string_length (''c.andi''))))) of s17170 => (case ((spc_matches_prefix0 s17170)) of Some ((_, s17180)) => (case ((string_drop s17170 s17180)) of s17190 => (case ((creg_name_matches_prefix s17190 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s17200)) => (case ((string_drop s17190 s17200)) of s17210 => (case ((sep_matches_prefix s17210)) of Some ((_, s17220)) => (case ((string_drop s17210 s17220)) of s17230 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s17230 :: (( 6 Word.word * ii)) option)) of Some ((imm, s17240)) => (let p00 = (string_drop s17230 s17240) in if (((p00 = ('''')))) then Some (rsd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s17150 :: " string " \ \\val _s1702_ : string -> maybe ((mword ty3 * mword ty6))\\ definition s1702 :: \ string \((3)Word.word*(6)Word.word)option \ where \ s1702 s17030 = ( (let s17040 = s17030 in if ((string_startswith s17040 (''c.srai''))) then (case ((string_drop s17040 ((string_length (''c.srai''))))) of s17050 => (case ((spc_matches_prefix0 s17050)) of Some ((_, s17060)) => (case ((string_drop s17050 s17060)) of s17070 => (case ((creg_name_matches_prefix s17070 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s17080)) => (case ((string_drop s17070 s17080)) of s17090 => (case ((sep_matches_prefix s17090)) of Some ((_, s17100)) => (case ((string_drop s17090 s17100)) of s17110 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s17110 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s17120)) => (let p00 = (string_drop s17110 s17120) in if (((p00 = ('''')))) then Some (rsd, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s17030 :: " string " \ \\val _s1690_ : string -> maybe ((mword ty3 * mword ty6))\\ definition s1690 :: \ string \((3)Word.word*(6)Word.word)option \ where \ s1690 s16910 = ( (let s16920 = s16910 in if ((string_startswith s16920 (''c.srli''))) then (case ((string_drop s16920 ((string_length (''c.srli''))))) of s16930 => (case ((spc_matches_prefix0 s16930)) of Some ((_, s16940)) => (case ((string_drop s16930 s16940)) of s16950 => (case ((creg_name_matches_prefix s16950 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s16960)) => (case ((string_drop s16950 s16960)) of s16970 => (case ((sep_matches_prefix s16970)) of Some ((_, s16980)) => (case ((string_drop s16970 s16980)) of s16990 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s16990 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s17000)) => (let p00 = (string_drop s16990 s17000) in if (((p00 = ('''')))) then Some (rsd, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s16910 :: " string " \ \\val _s1678_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1678 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1678 s16790 = ( (let s16800 = s16790 in if ((string_startswith s16800 (''c.lui''))) then (case ((string_drop s16800 ((string_length (''c.lui''))))) of s16810 => (case ((spc_matches_prefix0 s16810)) of Some ((_, s16820)) => (case ((string_drop s16810 s16820)) of s16830 => (case ((reg_name_matches_prefix s16830 :: (( 5 Word.word * ii)) option)) of Some ((rd, s16840)) => (case ((string_drop s16830 s16840)) of s16850 => (case ((sep_matches_prefix s16850)) of Some ((_, s16860)) => (case ((string_drop s16850 s16860)) of s16870 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s16870 :: (( 6 Word.word * ii)) option)) of Some ((imm, s16880)) => (let p00 = (string_drop s16870 s16880) in if (((p00 = ('''')))) then Some (rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s16790 :: " string " \ \\val _s1670_ : string -> maybe (mword ty6)\\ definition s1670 :: \ string \((6)Word.word)option \ where \ s1670 s16710 = ( (let s16720 = s16710 in if ((string_startswith s16720 (''c.addi16sp''))) then (case ((string_drop s16720 ((string_length (''c.addi16sp''))))) of s16730 => (case ((spc_matches_prefix0 s16730)) of Some ((_, s16740)) => (case ((string_drop s16730 s16740)) of s16750 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s16750 :: (( 6 Word.word * ii)) option)) of Some ((imm, s16760)) => (let p00 = (string_drop s16750 s16760) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) | _ => None ) ) else None))\ for s16710 :: " string " \ \\val _s1658_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1658 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1658 s16590 = ( (let s16600 = s16590 in if ((string_startswith s16600 (''c.li''))) then (case ((string_drop s16600 ((string_length (''c.li''))))) of s16610 => (case ((spc_matches_prefix0 s16610)) of Some ((_, s16620)) => (case ((string_drop s16610 s16620)) of s16630 => (case ((reg_name_matches_prefix s16630 :: (( 5 Word.word * ii)) option)) of Some ((rd, s16640)) => (case ((string_drop s16630 s16640)) of s16650 => (case ((sep_matches_prefix s16650)) of Some ((_, s16660)) => (case ((string_drop s16650 s16660)) of s16670 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s16670 :: (( 6 Word.word * ii)) option)) of Some ((imm, s16680)) => (let p00 = (string_drop s16670 s16680) in if (((p00 = ('''')))) then Some (rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s16590 :: " string " \ \\val _s1646_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1646 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1646 s16470 = ( (let s16480 = s16470 in if ((string_startswith s16480 (''c.addiw''))) then (case ((string_drop s16480 ((string_length (''c.addiw''))))) of s16490 => (case ((spc_matches_prefix0 s16490)) of Some ((_, s16500)) => (case ((string_drop s16490 s16500)) of s16510 => (case ((reg_name_matches_prefix s16510 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s16520)) => (case ((string_drop s16510 s16520)) of s16530 => (case ((sep_matches_prefix s16530)) of Some ((_, s16540)) => (case ((string_drop s16530 s16540)) of s16550 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s16550 :: (( 6 Word.word * ii)) option)) of Some ((imm, s16560)) => (let p00 = (string_drop s16550 s16560) in if (((p00 = ('''')))) then Some (rsd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s16470 :: " string " \ \\val _s1638_ : string -> maybe (mword ty11)\\ definition s1638 :: \ string \((11)Word.word)option \ where \ s1638 s16390 = ( (let s16400 = s16390 in if ((string_startswith s16400 (''c.jal''))) then (case ((string_drop s16400 ((string_length (''c.jal''))))) of s16410 => (case ((spc_matches_prefix0 s16410)) of Some ((_, s16420)) => (case ((string_drop s16410 s16420)) of s16430 => (case ((hex_bits_12_matches_prefix0 s16430 :: (( 12 Word.word * ii)) option)) of Some ((v__1184, s16440)) => if (((((subrange_vec_dec v__1184 (( 0 :: int):: ii) (( 0 :: int):: ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then (let (imm :: 11 Word.word) = ((subrange_vec_dec v__1184 (( 11 :: int):: ii) (( 1 :: int):: ii) :: 11 Word.word)) in (let (imm :: 11 Word.word) = ((subrange_vec_dec v__1184 (( 11 :: int):: ii) (( 1 :: int):: ii) :: 11 Word.word)) in (let p00 = (string_drop s16430 s16440) in if (((p00 = ('''')))) then Some imm else None))) else None | _ => None ) ) | _ => None ) ) else None))\ for s16390 :: " string " \ \\val _s1626_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s1626 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s1626 s16270 = ( (let s16280 = s16270 in if ((string_startswith s16280 (''c.addi''))) then (case ((string_drop s16280 ((string_length (''c.addi''))))) of s16290 => (case ((spc_matches_prefix0 s16290)) of Some ((_, s16300)) => (case ((string_drop s16290 s16300)) of s16310 => (case ((reg_name_matches_prefix s16310 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s16320)) => (case ((string_drop s16310 s16320)) of s16330 => (case ((sep_matches_prefix s16330)) of Some ((_, s16340)) => (case ((string_drop s16330 s16340)) of s16350 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s16350 :: (( 6 Word.word * ii)) option)) of Some ((nzi, s16360)) => (let p00 = (string_drop s16350 s16360) in if (((p00 = ('''')))) then Some (rsd, nzi) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s16270 :: " string " \ \\val _s1610_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s1610 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s1610 s16110 = ( (let s16120 = s16110 in if ((string_startswith s16120 (''c.sd''))) then (case ((string_drop s16120 ((string_length (''c.sd''))))) of s16130 => (case ((spc_matches_prefix0 s16130)) of Some ((_, s16140)) => (case ((string_drop s16130 s16140)) of s16150 => (case ((creg_name_matches_prefix s16150 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s16160)) => (case ((string_drop s16150 s16160)) of s16170 => (case ((sep_matches_prefix s16170)) of Some ((_, s16180)) => (case ((string_drop s16170 s16180)) of s16190 => (case ((creg_name_matches_prefix s16190 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s16200)) => (case ((string_drop s16190 s16200)) of s16210 => (case ((sep_matches_prefix s16210)) of Some ((_, s16220)) => (case ((string_drop s16210 s16220)) of s16230 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s16230 :: (( 8 Word.word * ii)) option)) of Some ((v__1186, s16240)) => if (((((subrange_vec_dec v__1186 (( 2 :: int):: ii) (( 0 :: int):: ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1186 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1186 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s16230 s16240) in if (((p00 = ('''')))) then Some (rsc1, rsc2, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s16110 :: " string " \ \\val _s1594_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s1594 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s1594 s15950 = ( (let s15960 = s15950 in if ((string_startswith s15960 (''c.sw''))) then (case ((string_drop s15960 ((string_length (''c.sw''))))) of s15970 => (case ((spc_matches_prefix0 s15970)) of Some ((_, s15980)) => (case ((string_drop s15970 s15980)) of s15990 => (case ((creg_name_matches_prefix s15990 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s16000)) => (case ((string_drop s15990 s16000)) of s16010 => (case ((sep_matches_prefix s16010)) of Some ((_, s16020)) => (case ((string_drop s16010 s16020)) of s16030 => (case ((creg_name_matches_prefix s16030 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s16040)) => (case ((string_drop s16030 s16040)) of s16050 => (case ((sep_matches_prefix s16050)) of Some ((_, s16060)) => (case ((string_drop s16050 s16060)) of s16070 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s16070 :: (( 7 Word.word * ii)) option)) of Some ((v__1188, s16080)) => if (((((subrange_vec_dec v__1188 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1188 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1188 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s16070 s16080) in if (((p00 = ('''')))) then Some (rsc1, rsc2, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s15950 :: " string " \ \\val _s1578_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s1578 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s1578 s15790 = ( (let s15800 = s15790 in if ((string_startswith s15800 (''c.ld''))) then (case ((string_drop s15800 ((string_length (''c.ld''))))) of s15810 => (case ((spc_matches_prefix0 s15810)) of Some ((_, s15820)) => (case ((string_drop s15810 s15820)) of s15830 => (case ((creg_name_matches_prefix s15830 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s15840)) => (case ((string_drop s15830 s15840)) of s15850 => (case ((sep_matches_prefix s15850)) of Some ((_, s15860)) => (case ((string_drop s15850 s15860)) of s15870 => (case ((creg_name_matches_prefix s15870 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s15880)) => (case ((string_drop s15870 s15880)) of s15890 => (case ((sep_matches_prefix s15890)) of Some ((_, s15900)) => (case ((string_drop s15890 s15900)) of s15910 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s15910 :: (( 8 Word.word * ii)) option)) of Some ((v__1190, s15920)) => if (((((subrange_vec_dec v__1190 (( 2 :: int):: ii) (( 0 :: int):: ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1190 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1190 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s15910 s15920) in if (((p00 = ('''')))) then Some (rdc, rsc, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s15790 :: " string " \ \\val _s1562_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s1562 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s1562 s15630 = ( (let s15640 = s15630 in if ((string_startswith s15640 (''c.lw''))) then (case ((string_drop s15640 ((string_length (''c.lw''))))) of s15650 => (case ((spc_matches_prefix0 s15650)) of Some ((_, s15660)) => (case ((string_drop s15650 s15660)) of s15670 => (case ((creg_name_matches_prefix s15670 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s15680)) => (case ((string_drop s15670 s15680)) of s15690 => (case ((sep_matches_prefix s15690)) of Some ((_, s15700)) => (case ((string_drop s15690 s15700)) of s15710 => (case ((creg_name_matches_prefix s15710 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s15720)) => (case ((string_drop s15710 s15720)) of s15730 => (case ((sep_matches_prefix s15730)) of Some ((_, s15740)) => (case ((string_drop s15730 s15740)) of s15750 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s15750 :: (( 7 Word.word * ii)) option)) of Some ((v__1192, s15760)) => if (((((subrange_vec_dec v__1192 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1192 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1192 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s15750 s15760) in if (((p00 = ('''')))) then Some (rdc, rsc, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s15630 :: " string " \ \\val _s1550_ : string -> maybe ((mword ty3 * mword ty8))\\ definition s1550 :: \ string \((3)Word.word*(8)Word.word)option \ where \ s1550 s15510 = ( (let s15520 = s15510 in if ((string_startswith s15520 (''c.addi4spn''))) then (case ((string_drop s15520 ((string_length (''c.addi4spn''))))) of s15530 => (case ((spc_matches_prefix0 s15530)) of Some ((_, s15540)) => (case ((string_drop s15530 s15540)) of s15550 => (case ((creg_name_matches_prefix s15550 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s15560)) => (case ((string_drop s15550 s15560)) of s15570 => (case ((sep_matches_prefix s15570)) of Some ((_, s15580)) => (case ((string_drop s15570 s15580)) of s15590 => (case ((hex_bits_10_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s15590 :: (( 10 Word.word * ii)) option)) of Some ((v__1194, s15600)) => if (((((subrange_vec_dec v__1194 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (nzimm :: 8 Word.word) = ((subrange_vec_dec v__1194 (( 9 :: int):: ii) (( 2 :: int):: ii) :: 8 Word.word)) in (let (nzimm :: 8 Word.word) = ((subrange_vec_dec v__1194 (( 9 :: int):: ii) (( 2 :: int):: ii) :: 8 Word.word)) in (let p00 = (string_drop s15590 s15600) in if (((p00 = ('''')))) then Some (rdc, nzimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s15510 :: " string " \ \\val _s1524_ : string -> maybe ((amoop * word_width * bool * bool * mword ty5 * mword ty5 * mword ty5))\\ definition s1524 :: \ string \(amoop*word_width*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1524 s15260 = ( (case ((amo_mnemonic_matches_prefix s15260)) of Some ((op1, s15270)) => (let s15280 = (string_drop s15260 s15270) in if ((string_startswith s15280 (''.''))) then (case ((string_drop s15280 ((string_length (''.''))))) of s15290 => (case ((size_mnemonic_matches_prefix s15290)) of Some ((width, s15300)) => (case ((string_drop s15290 s15300)) of s15310 => (case ((maybe_aq_matches_prefix s15310)) of Some ((aq, s15320)) => (case ((string_drop s15310 s15320)) of s15330 => (case ((maybe_rl_matches_prefix s15330)) of Some ((rl, s15340)) => (case ((string_drop s15330 s15340)) of s15350 => (case ((spc_matches_prefix0 s15350)) of Some ((_, s15360)) => (case ((string_drop s15350 s15360)) of s15370 => (case ((reg_name_matches_prefix s15370 :: (( 5 Word.word * ii)) option)) of Some ((rd, s15380)) => (case ((string_drop s15370 s15380)) of s15390 => (case ((sep_matches_prefix s15390)) of Some ((_, s15400)) => (case ((string_drop s15390 s15400)) of s15410 => (case ((reg_name_matches_prefix s15410 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s15420)) => (case ((string_drop s15410 s15420)) of s15430 => (case ((sep_matches_prefix s15430)) of Some ((_, s15440)) => (let s15450 = (string_drop s15430 s15440) in if ((string_startswith s15450 (''(''))) then (case ((string_drop s15450 ((string_length (''(''))))) of s15460 => (case ((reg_name_matches_prefix s15460 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s15470)) => (let s15480 = (string_drop s15460 s15470) in if ((string_startswith s15480 ('')''))) then (let p00 = (string_drop s15480 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (op1, width, aq, rl, rd, rs2, rs1) else None) else None) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ))\ for s15260 :: " string " \ \\val _s1502_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty5 * mword ty5))\\ definition s1502 :: \ string \(word_width*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1502 s15030 = ( (let s15040 = s15030 in if ((string_startswith s15040 (''sc.''))) then (case ((string_drop s15040 ((string_length (''sc.''))))) of s15050 => (case ((size_mnemonic_matches_prefix s15050)) of Some ((size1, s15060)) => (case ((string_drop s15050 s15060)) of s15070 => (case ((maybe_aq_matches_prefix s15070)) of Some ((aq, s15080)) => (case ((string_drop s15070 s15080)) of s15090 => (case ((maybe_rl_matches_prefix s15090)) of Some ((rl, s15100)) => (case ((string_drop s15090 s15100)) of s15110 => (case ((spc_matches_prefix0 s15110)) of Some ((_, s15120)) => (case ((string_drop s15110 s15120)) of s15130 => (case ((reg_name_matches_prefix s15130 :: (( 5 Word.word * ii)) option)) of Some ((rd, s15140)) => (case ((string_drop s15130 s15140)) of s15150 => (case ((sep_matches_prefix s15150)) of Some ((_, s15160)) => (case ((string_drop s15150 s15160)) of s15170 => (case ((reg_name_matches_prefix s15170 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s15180)) => (case ((string_drop s15170 s15180)) of s15190 => (case ((sep_matches_prefix s15190)) of Some ((_, s15200)) => (case ((string_drop s15190 s15200)) of s15210 => (case ((reg_name_matches_prefix s15210 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s15220)) => (let p00 = (string_drop s15210 s15220) in if (((p00 = ('''')))) then Some (size1, aq, rl, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s15030 :: " string " \ \\val _s1484_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty5))\\ definition s1484 :: \ string \(word_width*bool*bool*(5)Word.word*(5)Word.word)option \ where \ s1484 s14850 = ( (let s14860 = s14850 in if ((string_startswith s14860 (''lr.''))) then (case ((string_drop s14860 ((string_length (''lr.''))))) of s14870 => (case ((size_mnemonic_matches_prefix s14870)) of Some ((size1, s14880)) => (case ((string_drop s14870 s14880)) of s14890 => (case ((maybe_aq_matches_prefix s14890)) of Some ((aq, s14900)) => (case ((string_drop s14890 s14900)) of s14910 => (case ((maybe_rl_matches_prefix s14910)) of Some ((rl, s14920)) => (case ((string_drop s14910 s14920)) of s14930 => (case ((spc_matches_prefix0 s14930)) of Some ((_, s14940)) => (case ((string_drop s14930 s14940)) of s14950 => (case ((reg_name_matches_prefix s14950 :: (( 5 Word.word * ii)) option)) of Some ((rd, s14960)) => (case ((string_drop s14950 s14960)) of s14970 => (case ((sep_matches_prefix s14970)) of Some ((_, s14980)) => (case ((string_drop s14970 s14980)) of s14990 => (case ((reg_name_matches_prefix s14990 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s15000)) => (let p00 = (string_drop s14990 s15000) in if (((p00 = ('''')))) then Some (size1, aq, rl, rd, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s14850 :: " string " \ \\val _s1472_ : string -> maybe ((mword ty5 * mword ty5))\\ definition s1472 :: \ string \((5)Word.word*(5)Word.word)option \ where \ s1472 s14730 = ( (let s14740 = s14730 in if ((string_startswith s14740 (''sfence.vma''))) then (case ((string_drop s14740 ((string_length (''sfence.vma''))))) of s14750 => (case ((spc_matches_prefix0 s14750)) of Some ((_, s14760)) => (case ((string_drop s14750 s14760)) of s14770 => (case ((reg_name_matches_prefix s14770 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s14780)) => (case ((string_drop s14770 s14780)) of s14790 => (case ((sep_matches_prefix s14790)) of Some ((_, s14800)) => (case ((string_drop s14790 s14800)) of s14810 => (case ((reg_name_matches_prefix s14810 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s14820)) => (let p00 = (string_drop s14810 s14820) in if (((p00 = ('''')))) then Some (rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s14730 :: " string " \ \\val _s1460_ : string -> maybe ((mword ty4 * mword ty4))\\ definition s1460 :: \ string \((4)Word.word*(4)Word.word)option \ where \ s1460 s14610 = ( (let s14620 = s14610 in if ((string_startswith s14620 (''fence.tso''))) then (case ((string_drop s14620 ((string_length (''fence.tso''))))) of s14630 => (case ((spc_matches_prefix0 s14630)) of Some ((_, s14640)) => (case ((string_drop s14630 s14640)) of s14650 => (case ((fence_bits_matches_prefix s14650 :: (( 4 Word.word * ii)) option)) of Some ((pred, s14660)) => (case ((string_drop s14650 s14660)) of s14670 => (case ((sep_matches_prefix s14670)) of Some ((_, s14680)) => (case ((string_drop s14670 s14680)) of s14690 => (case ((fence_bits_matches_prefix s14690 :: (( 4 Word.word * ii)) option)) of Some ((succ, s14700)) => (let p00 = (string_drop s14690 s14700) in if (((p00 = ('''')))) then Some (pred, succ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s14610 :: " string " \ \\val _s1448_ : string -> maybe ((mword ty4 * mword ty4))\\ definition s1448 :: \ string \((4)Word.word*(4)Word.word)option \ where \ s1448 s14490 = ( (let s14500 = s14490 in if ((string_startswith s14500 (''fence''))) then (case ((string_drop s14500 ((string_length (''fence''))))) of s14510 => (case ((spc_matches_prefix0 s14510)) of Some ((_, s14520)) => (case ((string_drop s14510 s14520)) of s14530 => (case ((fence_bits_matches_prefix s14530 :: (( 4 Word.word * ii)) option)) of Some ((pred, s14540)) => (case ((string_drop s14530 s14540)) of s14550 => (case ((sep_matches_prefix s14550)) of Some ((_, s14560)) => (case ((string_drop s14550 s14560)) of s14570 => (case ((fence_bits_matches_prefix s14570 :: (( 4 Word.word * ii)) option)) of Some ((succ, s14580)) => (let p00 = (string_drop s14570 s14580) in if (((p00 = ('''')))) then Some (pred, succ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s14490 :: " string " \ \\val _s1431_ : string -> maybe ((sopw * mword ty5 * mword ty5 * mword ty5))\\ definition s1431 :: \ string \(sopw*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1431 s14330 = ( (case ((shiftiwop_mnemonic_matches_prefix s14330)) of Some ((op1, s14340)) => (case ((string_drop s14330 s14340)) of s14350 => (case ((spc_matches_prefix0 s14350)) of Some ((_, s14360)) => (case ((string_drop s14350 s14360)) of s14370 => (case ((reg_name_matches_prefix s14370 :: (( 5 Word.word * ii)) option)) of Some ((rd, s14380)) => (case ((string_drop s14370 s14380)) of s14390 => (case ((sep_matches_prefix s14390)) of Some ((_, s14400)) => (case ((string_drop s14390 s14400)) of s14410 => (case ((reg_name_matches_prefix s14410 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s14420)) => (case ((string_drop s14410 s14420)) of s14430 => (case ((sep_matches_prefix s14430)) of Some ((_, s14440)) => (case ((string_drop s14430 s14440)) of s14450 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s14450 :: (( 5 Word.word * ii)) option)) of Some ((shamt, s14460)) => (let p00 = (string_drop s14450 s14460) in if (((p00 = ('''')))) then Some (op1, rd, rs1, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s14330 :: " string " \ \\val _s1414_ : string -> maybe ((ropw * mword ty5 * mword ty5 * mword ty5))\\ definition s1414 :: \ string \(ropw*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1414 s14160 = ( (case ((rtypew_mnemonic_matches_prefix s14160)) of Some ((op1, s14170)) => (case ((string_drop s14160 s14170)) of s14180 => (case ((spc_matches_prefix0 s14180)) of Some ((_, s14190)) => (case ((string_drop s14180 s14190)) of s14200 => (case ((reg_name_matches_prefix s14200 :: (( 5 Word.word * ii)) option)) of Some ((rd, s14210)) => (case ((string_drop s14200 s14210)) of s14220 => (case ((sep_matches_prefix s14220)) of Some ((_, s14230)) => (case ((string_drop s14220 s14230)) of s14240 => (case ((reg_name_matches_prefix s14240 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s14250)) => (case ((string_drop s14240 s14250)) of s14260 => (case ((sep_matches_prefix s14260)) of Some ((_, s14270)) => (case ((string_drop s14260 s14270)) of s14280 => (case ((reg_name_matches_prefix s14280 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s14290)) => (let p00 = (string_drop s14280 s14290) in if (((p00 = ('''')))) then Some (op1, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s14160 :: " string " \ \\val _s1397_ : string -> maybe ((sop * mword ty5 * mword ty5 * mword ty5))\\ definition s1397 :: \ string \(sop*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1397 s13990 = ( (case ((shiftw_mnemonic_matches_prefix s13990)) of Some ((op1, s14000)) => (case ((string_drop s13990 s14000)) of s14010 => (case ((spc_matches_prefix0 s14010)) of Some ((_, s14020)) => (case ((string_drop s14010 s14020)) of s14030 => (case ((reg_name_matches_prefix s14030 :: (( 5 Word.word * ii)) option)) of Some ((rd, s14040)) => (case ((string_drop s14030 s14040)) of s14050 => (case ((sep_matches_prefix s14050)) of Some ((_, s14060)) => (case ((string_drop s14050 s14060)) of s14070 => (case ((reg_name_matches_prefix s14070 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s14080)) => (case ((string_drop s14070 s14080)) of s14090 => (case ((sep_matches_prefix s14090)) of Some ((_, s14100)) => (case ((string_drop s14090 s14100)) of s14110 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s14110 :: (( 5 Word.word * ii)) option)) of Some ((shamt, s14120)) => (let p00 = (string_drop s14110 s14120) in if (((p00 = ('''')))) then Some (op1, rd, rs1, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s13990 :: " string " \ \\val _s1381_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12))\\ definition s1381 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word)option \ where \ s1381 s13820 = ( (let s13830 = s13820 in if ((string_startswith s13830 (''addiw''))) then (case ((string_drop s13830 ((string_length (''addiw''))))) of s13840 => (case ((spc_matches_prefix0 s13840)) of Some ((_, s13850)) => (case ((string_drop s13840 s13850)) of s13860 => (case ((reg_name_matches_prefix s13860 :: (( 5 Word.word * ii)) option)) of Some ((rd, s13870)) => (case ((string_drop s13860 s13870)) of s13880 => (case ((sep_matches_prefix s13880)) of Some ((_, s13890)) => (case ((string_drop s13880 s13890)) of s13900 => (case ((reg_name_matches_prefix s13900 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s13910)) => (case ((string_drop s13900 s13910)) of s13920 => (case ((sep_matches_prefix s13920)) of Some ((_, s13930)) => (case ((string_drop s13920 s13930)) of s13940 => (case ((hex_bits_12_matches_prefix0 s13940 :: (( 12 Word.word * ii)) option)) of Some ((imm, s13950)) => (let p00 = (string_drop s13940 s13950) in if (((p00 = ('''')))) then Some (rd, rs1, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s13820 :: " string " \ \\val _s1353_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty12 * mword ty5))\\ definition s1353 :: \ string \(word_width*bool*bool*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s1353 s13540 = ( (let s13550 = s13540 in if ((string_startswith s13550 (''s''))) then (case ((string_drop s13550 ((string_length (''s''))))) of s13560 => (case ((size_mnemonic_matches_prefix s13560)) of Some ((size1, s13570)) => (case ((string_drop s13560 s13570)) of s13580 => (case ((maybe_aq_matches_prefix s13580)) of Some ((aq, s13590)) => (case ((string_drop s13580 s13590)) of s13600 => (case ((maybe_rl_matches_prefix s13600)) of Some ((rl, s13610)) => (case ((string_drop s13600 s13610)) of s13620 => (case ((spc_matches_prefix0 s13620)) of Some ((_, s13630)) => (case ((string_drop s13620 s13630)) of s13640 => (case ((reg_name_matches_prefix s13640 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s13650)) => (case ((string_drop s13640 s13650)) of s13660 => (case ((sep_matches_prefix s13660)) of Some ((_, s13670)) => (case ((string_drop s13660 s13670)) of s13680 => (case ((hex_bits_12_matches_prefix0 s13680 :: (( 12 Word.word * ii)) option)) of Some ((imm, s13690)) => (case ((string_drop s13680 s13690)) of s13700 => (case ((opt_spc_matches_prefix0 s13700)) of Some ((_, s13710)) => (let s13720 = (string_drop s13700 s13710) in if ((string_startswith s13720 (''(''))) then (case ((string_drop s13720 ((string_length (''(''))))) of s13730 => (case ((opt_spc_matches_prefix0 s13730)) of Some ((_, s13740)) => (case ((string_drop s13730 s13740)) of s13750 => (case ((reg_name_matches_prefix s13750 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s13760)) => (case ((string_drop s13750 s13760)) of s13770 => (case ((opt_spc_matches_prefix0 s13770)) of Some ((_, s13780)) => (let s13790 = (string_drop s13770 s13780) in if ((string_startswith s13790 ('')''))) then (let p00 = (string_drop s13790 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (size1, aq, rl, rs2, imm, rs1) else None) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s13540 :: " string " \ \\val _s1323_ : string -> maybe ((word_width * bool * bool * bool * mword ty5 * mword ty12 * mword ty5))\\ definition s1323 :: \ string \(word_width*bool*bool*bool*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s1323 s13240 = ( (let s13250 = s13240 in if ((string_startswith s13250 (''l''))) then (case ((string_drop s13250 ((string_length (''l''))))) of s13260 => (case ((size_mnemonic_matches_prefix s13260)) of Some ((size1, s13270)) => (case ((string_drop s13260 s13270)) of s13280 => (case ((maybe_u_matches_prefix s13280)) of Some ((is_unsigned, s13290)) => (case ((string_drop s13280 s13290)) of s13300 => (case ((maybe_aq_matches_prefix s13300)) of Some ((aq, s13310)) => (case ((string_drop s13300 s13310)) of s13320 => (case ((maybe_rl_matches_prefix s13320)) of Some ((rl, s13330)) => (case ((string_drop s13320 s13330)) of s13340 => (case ((spc_matches_prefix0 s13340)) of Some ((_, s13350)) => (case ((string_drop s13340 s13350)) of s13360 => (case ((reg_name_matches_prefix s13360 :: (( 5 Word.word * ii)) option)) of Some ((rd, s13370)) => (case ((string_drop s13360 s13370)) of s13380 => (case ((sep_matches_prefix s13380)) of Some ((_, s13390)) => (case ((string_drop s13380 s13390)) of s13400 => (case ((hex_bits_12_matches_prefix0 s13400 :: (( 12 Word.word * ii)) option)) of Some ((imm, s13410)) => (case ((string_drop s13400 s13410)) of s13420 => (case ((opt_spc_matches_prefix0 s13420)) of Some ((_, s13430)) => (let s13440 = (string_drop s13420 s13430) in if ((string_startswith s13440 (''(''))) then (case ((string_drop s13440 ((string_length (''(''))))) of s13450 => (case ((opt_spc_matches_prefix0 s13450)) of Some ((_, s13460)) => (case ((string_drop s13450 s13460)) of s13470 => (case ((reg_name_matches_prefix s13470 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s13480)) => (case ((string_drop s13470 s13480)) of s13490 => (case ((opt_spc_matches_prefix0 s13490)) of Some ((_, s13500)) => (let s13510 = (string_drop s13490 s13500) in if ((string_startswith s13510 ('')''))) then (let p00 = (string_drop s13510 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (size1, is_unsigned, aq, rl, rd, imm, rs1) else None) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s13240 :: " string " \ \\val _s1306_ : string -> maybe ((rop * mword ty5 * mword ty5 * mword ty5))\\ definition s1306 :: \ string \(rop*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s1306 s13080 = ( (case ((rtype_mnemonic_matches_prefix s13080)) of Some ((op1, s13090)) => (case ((string_drop s13080 s13090)) of s13100 => (case ((spc_matches_prefix0 s13100)) of Some ((_, s13110)) => (case ((string_drop s13100 s13110)) of s13120 => (case ((reg_name_matches_prefix s13120 :: (( 5 Word.word * ii)) option)) of Some ((rd, s13130)) => (case ((string_drop s13120 s13130)) of s13140 => (case ((sep_matches_prefix s13140)) of Some ((_, s13150)) => (case ((string_drop s13140 s13150)) of s13160 => (case ((reg_name_matches_prefix s13160 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s13170)) => (case ((string_drop s13160 s13170)) of s13180 => (case ((sep_matches_prefix s13180)) of Some ((_, s13190)) => (case ((string_drop s13180 s13190)) of s13200 => (case ((reg_name_matches_prefix s13200 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s13210)) => (let p00 = (string_drop s13200 s13210) in if (((p00 = ('''')))) then Some (op1, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s13080 :: " string " \ \\val _s1289_ : string -> maybe ((sop * mword ty5 * mword ty5 * mword ty6))\\ definition s1289 :: \ string \(sop*(5)Word.word*(5)Word.word*(6)Word.word)option \ where \ s1289 s12910 = ( (case ((shiftiop_mnemonic_matches_prefix s12910)) of Some ((op1, s12920)) => (case ((string_drop s12910 s12920)) of s12930 => (case ((spc_matches_prefix0 s12930)) of Some ((_, s12940)) => (case ((string_drop s12930 s12940)) of s12950 => (case ((reg_name_matches_prefix s12950 :: (( 5 Word.word * ii)) option)) of Some ((rd, s12960)) => (case ((string_drop s12950 s12960)) of s12970 => (case ((sep_matches_prefix s12970)) of Some ((_, s12980)) => (case ((string_drop s12970 s12980)) of s12990 => (case ((reg_name_matches_prefix s12990 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s13000)) => (case ((string_drop s12990 s13000)) of s13010 => (case ((sep_matches_prefix s13010)) of Some ((_, s13020)) => (case ((string_drop s13010 s13020)) of s13030 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s13030 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s13040)) => (let p00 = (string_drop s13030 s13040) in if (((p00 = ('''')))) then Some (op1, rd, rs1, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s12910 :: " string " \ \\val _s1272_ : string -> maybe ((iop * mword ty5 * mword ty5 * mword ty12))\\ definition s1272 :: \ string \(iop*(5)Word.word*(5)Word.word*(12)Word.word)option \ where \ s1272 s12740 = ( (case ((itype_mnemonic_matches_prefix s12740)) of Some ((op1, s12750)) => (case ((string_drop s12740 s12750)) of s12760 => (case ((spc_matches_prefix0 s12760)) of Some ((_, s12770)) => (case ((string_drop s12760 s12770)) of s12780 => (case ((reg_name_matches_prefix s12780 :: (( 5 Word.word * ii)) option)) of Some ((rd, s12790)) => (case ((string_drop s12780 s12790)) of s12800 => (case ((sep_matches_prefix s12800)) of Some ((_, s12810)) => (case ((string_drop s12800 s12810)) of s12820 => (case ((reg_name_matches_prefix s12820 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s12830)) => (case ((string_drop s12820 s12830)) of s12840 => (case ((sep_matches_prefix s12840)) of Some ((_, s12850)) => (case ((string_drop s12840 s12850)) of s12860 => (case ((hex_bits_12_matches_prefix0 s12860 :: (( 12 Word.word * ii)) option)) of Some ((imm, s12870)) => (let p00 = (string_drop s12860 s12870) in if (((p00 = ('''')))) then Some (op1, rd, rs1, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s12740 :: " string " \ \\val _s1255_ : string -> maybe ((bop * mword ty5 * mword ty5 * mword ty13))\\ definition s1255 :: \ string \(bop*(5)Word.word*(5)Word.word*(13)Word.word)option \ where \ s1255 s12570 = ( (case ((btype_mnemonic_matches_prefix s12570)) of Some ((op1, s12580)) => (case ((string_drop s12570 s12580)) of s12590 => (case ((spc_matches_prefix0 s12590)) of Some ((_, s12600)) => (case ((string_drop s12590 s12600)) of s12610 => (case ((reg_name_matches_prefix s12610 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s12620)) => (case ((string_drop s12610 s12620)) of s12630 => (case ((sep_matches_prefix s12630)) of Some ((_, s12640)) => (case ((string_drop s12630 s12640)) of s12650 => (case ((reg_name_matches_prefix s12650 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s12660)) => (case ((string_drop s12650 s12660)) of s12670 => (case ((sep_matches_prefix s12670)) of Some ((_, s12680)) => (case ((string_drop s12670 s12680)) of s12690 => (case ((hex_bits_13_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s12690 :: (( 13 Word.word * ii)) option)) of Some ((imm, s12700)) => (let p00 = (string_drop s12690 s12700) in if (((p00 = ('''')))) then Some (op1, rs1, rs2, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s12570 :: " string " \ \\val _s1239_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12))\\ definition s1239 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word)option \ where \ s1239 s12400 = ( (let s12410 = s12400 in if ((string_startswith s12410 (''jalr''))) then (case ((string_drop s12410 ((string_length (''jalr''))))) of s12420 => (case ((spc_matches_prefix0 s12420)) of Some ((_, s12430)) => (case ((string_drop s12420 s12430)) of s12440 => (case ((reg_name_matches_prefix s12440 :: (( 5 Word.word * ii)) option)) of Some ((rd, s12450)) => (case ((string_drop s12440 s12450)) of s12460 => (case ((sep_matches_prefix s12460)) of Some ((_, s12470)) => (case ((string_drop s12460 s12470)) of s12480 => (case ((reg_name_matches_prefix s12480 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s12490)) => (case ((string_drop s12480 s12490)) of s12500 => (case ((sep_matches_prefix s12500)) of Some ((_, s12510)) => (case ((string_drop s12500 s12510)) of s12520 => (case ((hex_bits_12_matches_prefix0 s12520 :: (( 12 Word.word * ii)) option)) of Some ((imm, s12530)) => (let p00 = (string_drop s12520 s12530) in if (((p00 = ('''')))) then Some (rd, rs1, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s12400 :: " string " \ \\val _s1227_ : string -> maybe ((mword ty5 * mword ty21))\\ definition s1227 :: \ string \((5)Word.word*(21)Word.word)option \ where \ s1227 s12280 = ( (let s12290 = s12280 in if ((string_startswith s12290 (''jal''))) then (case ((string_drop s12290 ((string_length (''jal''))))) of s12300 => (case ((spc_matches_prefix0 s12300)) of Some ((_, s12310)) => (case ((string_drop s12300 s12310)) of s12320 => (case ((reg_name_matches_prefix s12320 :: (( 5 Word.word * ii)) option)) of Some ((rd, s12330)) => (case ((string_drop s12320 s12330)) of s12340 => (case ((sep_matches_prefix s12340)) of Some ((_, s12350)) => (case ((string_drop s12340 s12350)) of s12360 => (case ((hex_bits_21_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s12360 :: (( 21 Word.word * ii)) option)) of Some ((imm, s12370)) => (let p00 = (string_drop s12360 s12370) in if (((p00 = ('''')))) then Some (rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s12280 :: " string " \ \\val _s1214_ : string -> maybe ((uop * mword ty5 * mword ty20))\\ definition s1214 :: \ string \(uop*(5)Word.word*(20)Word.word)option \ where \ s1214 s12160 = ( (case ((utype_mnemonic_matches_prefix s12160)) of Some ((op1, s12170)) => (case ((string_drop s12160 s12170)) of s12180 => (case ((spc_matches_prefix0 s12180)) of Some ((_, s12190)) => (case ((string_drop s12180 s12190)) of s12200 => (case ((reg_name_matches_prefix s12200 :: (( 5 Word.word * ii)) option)) of Some ((rd, s12210)) => (case ((string_drop s12200 s12210)) of s12220 => (case ((sep_matches_prefix s12220)) of Some ((_, s12230)) => (case ((string_drop s12220 s12230)) of s12240 => (case ((hex_bits_20_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s12240 :: (( 20 Word.word * ii)) option)) of Some ((imm, s12250)) => (let p00 = (string_drop s12240 s12250) in if (((p00 = ('''')))) then Some (op1, rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s12160 :: " string " definition assembly_backwards :: \ string \((register_value),(ast),(exception))monad \ where \ assembly_backwards arg1 = ( (let s12260 = arg1 in if ((case ((s1214 s12260 :: ((uop * 5 Word.word * 20 Word.word))option)) of Some ((op1, rd, imm)) => True | _ => False )) then (case (s1214 s12260 :: (( uop * 5 Word.word * 20 Word.word)) option) of (Some ((op1, rd, imm))) => return (UTYPE (imm, rd, op1)) ) else if ((case ((s1227 s12260 :: (( 5 Word.word * 21 Word.word))option)) of Some ((rd, imm)) => True | _ => False )) then (case (s1227 s12260 :: (( 5 Word.word * 21 Word.word)) option) of (Some ((rd, imm))) => return (RISCV_JAL (imm, rd)) ) else if ((case ((s1239 s12260 :: (( 5 Word.word * 5 Word.word * 12 Word.word))option)) of Some ((rd, rs1, imm)) => True | _ => False )) then (case (s1239 s12260 :: (( 5 Word.word * 5 Word.word * 12 Word.word)) option) of (Some ((rd, rs1, imm))) => return (RISCV_JALR (imm, rs1, rd)) ) else if ((case ((s1255 s12260 :: ((bop * 5 Word.word * 5 Word.word * 13 Word.word))option)) of Some ((op1, rs1, rs2, imm)) => True | _ => False )) then (case (s1255 s12260 :: (( bop * 5 Word.word * 5 Word.word * 13 Word.word)) option) of (Some ((op1, rs1, rs2, imm))) => return (BTYPE (imm, rs2, rs1, op1)) ) else if ((case ((s1272 s12260 :: ((iop * 5 Word.word * 5 Word.word * 12 Word.word))option)) of Some ((op1, rd, rs1, imm)) => True | _ => False )) then (case (s1272 s12260 :: (( iop * 5 Word.word * 5 Word.word * 12 Word.word)) option) of (Some ((op1, rd, rs1, imm))) => return (ITYPE (imm, rs1, rd, op1)) ) else if ((case ((s1289 s12260 :: ((sop * 5 Word.word * 5 Word.word * 6 Word.word))option)) of Some ((op1, rd, rs1, shamt)) => True | _ => False )) then (case (s1289 s12260 :: (( sop * 5 Word.word * 5 Word.word * 6 Word.word)) option) of (Some ((op1, rd, rs1, shamt))) => return (SHIFTIOP (shamt, rs1, rd, op1)) ) else if ((case ((s1306 s12260 :: ((rop * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, rd, rs1, rs2)) => True | _ => False )) then (case (s1306 s12260 :: (( rop * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, rd, rs1, rs2))) => return (RTYPE (rs2, rs1, rd, op1)) ) else if ((case ((s1323 s12260 :: ((word_width * bool * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((size1, is_unsigned, aq, rl, rd, imm, rs1)) => True | _ => False )) then (case (s1323 s12260 :: (( word_width * bool * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((size1, is_unsigned, aq, rl, rd, imm, rs1))) => return (LOAD (imm, rs1, rd, is_unsigned, size1, aq, rl)) ) else if ((case ((s1353 s12260 :: ((word_width * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((size1, aq, rl, rs2, imm, rs1)) => True | _ => False )) then (case (s1353 s12260 :: (( word_width * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((size1, aq, rl, rs2, imm, rs1))) => return (STORE (imm, rs2, rs1, size1, aq, rl)) ) else if ((case ((s1381 s12260 :: (( 5 Word.word * 5 Word.word * 12 Word.word))option)) of Some ((rd, rs1, imm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1381 s12260 :: (( 5 Word.word * 5 Word.word * 12 Word.word)) option) of (Some ((rd, rs1, imm))) => return (ADDIW (imm, rs1, rd)) ) else if ((case ((s1397 s12260 :: ((sop * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, rd, rs1, shamt)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1397 s12260 :: (( sop * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, rd, rs1, shamt))) => return (SHIFTW (shamt, rs1, rd, op1)) ) else if ((case ((s1414 s12260 :: ((ropw * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, rd, rs1, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1414 s12260 :: (( ropw * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, rd, rs1, rs2))) => return (RTYPEW (rs2, rs1, rd, op1)) ) else if ((case ((s1431 s12260 :: ((sopw * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, rd, rs1, shamt)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1431 s12260 :: (( sopw * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, rd, rs1, shamt))) => return (SHIFTIWOP (shamt, rs1, rd, op1)) ) else if ((case ((s1448 s12260 :: (( 4 Word.word * 4 Word.word))option)) of Some ((pred, succ)) => True | _ => False )) then (case (s1448 s12260 :: (( 4 Word.word * 4 Word.word)) option) of (Some ((pred, succ))) => return (FENCE (pred, succ)) ) else if ((case ((s1460 s12260 :: (( 4 Word.word * 4 Word.word))option)) of Some ((pred, succ)) => True | _ => False )) then (case (s1460 s12260 :: (( 4 Word.word * 4 Word.word)) option) of (Some ((pred, succ))) => return (FENCE_TSO (pred, succ)) ) else if (((s12260 = (''fence.i'')))) then return (FENCEI () ) else if (((s12260 = (''ecall'')))) then return (ECALL () ) else if (((s12260 = (''mret'')))) then return (MRET () ) else if (((s12260 = (''sret'')))) then return (SRET () ) else if (((s12260 = (''ebreak'')))) then return (EBREAK () ) else if (((s12260 = (''wfi'')))) then return (WFI () ) else if ((case ((s1472 s12260 :: (( 5 Word.word * 5 Word.word))option)) of Some ((rs1, rs2)) => True | _ => False )) then (case (s1472 s12260 :: (( 5 Word.word * 5 Word.word)) option) of (Some ((rs1, rs2))) => return (SFENCE_VMA (rs1, rs2)) ) else if ((case ((s1484 s12260 :: ((word_width * bool * bool * 5 Word.word * 5 Word.word))option)) of Some ((size1, aq, rl, rd, rs1)) => True | _ => False )) then (case (s1484 s12260 :: (( word_width * bool * bool * 5 Word.word * 5 Word.word)) option) of (Some ((size1, aq, rl, rd, rs1))) => return (LOADRES (aq, rl, rs1, size1, rd)) ) else if ((case ((s1502 s12260 :: ((word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((size1, aq, rl, rd, rs1, rs2)) => True | _ => False )) then (case (s1502 s12260 :: (( word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((size1, aq, rl, rd, rs1, rs2))) => return (STORECON (aq, rl, rs2, rs1, size1, rd)) ) else if ((case ((s1524 s12260 :: ((amoop * word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, width, aq, rl, rd, rs2, rs1)) => True | _ => False )) then (case (s1524 s12260 :: (( amoop * word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, width, aq, rl, rd, rs2, rs1))) => return (AMO (op1, aq, rl, rs2, rs1, width, rd)) ) else if (((s12260 = (''c.nop'')))) then return (C_NOP () ) else if ((case ((s1550 s12260 :: (( 3 Word.word * 8 Word.word))option)) of Some ((rdc, nzimm)) => (nzimm \ ( 0x00 :: 8 Word.word)) | _ => False )) then (case (s1550 s12260 :: (( 3 Word.word * 8 Word.word)) option) of (Some ((rdc, nzimm))) => return (C_ADDI4SPN (rdc, nzimm)) ) else if ((case ((s1562 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rdc, rsc, uimm)) => True | _ => False )) then (case (s1562 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rdc, rsc, uimm))) => return (C_LW (uimm, rsc, rdc)) ) else if ((case ((s1578 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rdc, rsc, uimm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1578 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rdc, rsc, uimm))) => return (C_LD (uimm, rsc, rdc)) ) else if ((case ((s1594 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rsc1, rsc2, uimm)) => True | _ => False )) then (case (s1594 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rsc1, rsc2, uimm))) => return (C_SW (uimm, rsc1, rsc2)) ) else if ((case ((s1610 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rsc1, rsc2, uimm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1610 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rsc1, rsc2, uimm))) => return (C_SD (uimm, rsc1, rsc2)) ) else if ((case ((s1626 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rsd, nzi)) => ((((nzi \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))) | _ => False )) then (case (s1626 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rsd, nzi))) => return (C_ADDI (nzi, rsd)) ) else if ((case ((s1638 s12260 :: ( 11 Word.word)option)) of Some (imm) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s1638 s12260 :: ( 11 Word.word) option) of (Some (imm)) => return (C_JAL imm) ) else if ((case ((s1646 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rsd, imm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1646 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rsd, imm))) => return (C_ADDIW (imm, rsd)) ) else if ((case ((s1658 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, imm)) => (rd \ zreg) | _ => False )) then (case (s1658 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, imm))) => return (C_LI (imm, rd)) ) else if ((case ((s1670 s12260 :: ( 6 Word.word)option)) of Some (imm) => (imm \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s1670 s12260 :: ( 6 Word.word) option) of (Some (imm)) => return (C_ADDI16SP imm) ) else if ((case ((s1678 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, imm)) => ((((rd \ zreg))) \ ((((((rd \ sp))) \ (((imm \ ( 0b000000 :: 6 Word.word)))))))) | _ => False )) then (case (s1678 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, imm))) => return (C_LUI (imm, rd)) ) else if ((case ((s1690 s12260 :: (( 3 Word.word * 6 Word.word))option)) of Some ((rsd, shamt)) => (shamt \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s1690 s12260 :: (( 3 Word.word * 6 Word.word)) option) of (Some ((rsd, shamt))) => return (C_SRLI (shamt, rsd)) ) else if ((case ((s1702 s12260 :: (( 3 Word.word * 6 Word.word))option)) of Some ((rsd, shamt)) => (shamt \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s1702 s12260 :: (( 3 Word.word * 6 Word.word)) option) of (Some ((rsd, shamt))) => return (C_SRAI (shamt, rsd)) ) else if ((case ((s1714 s12260 :: (( 3 Word.word * 6 Word.word))option)) of Some ((rsd, imm)) => True | _ => False )) then (case (s1714 s12260 :: (( 3 Word.word * 6 Word.word)) option) of (Some ((rsd, imm))) => return (C_ANDI (imm, rsd)) ) else if ((case ((s1726 s12260 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => True | _ => False )) then (case (s1726 s12260 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => return (C_SUB (rsd, rs2)) ) else if ((case ((s1738 s12260 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => True | _ => False )) then (case (s1738 s12260 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => return (C_XOR (rsd, rs2)) ) else if ((case ((s1750 s12260 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => True | _ => False )) then (case (s1750 s12260 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => return (C_OR (rsd, rs2)) ) else if ((case ((s1762 s12260 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => True | _ => False )) then (case (s1762 s12260 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => return (C_AND (rsd, rs2)) ) else if ((case ((s1774 s12260 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1774 s12260 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => return (C_SUBW (rsd, rs2)) ) else if ((case ((s1786 s12260 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1786 s12260 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => return (C_ADDW (rsd, rs2)) ) else if ((case ((s1798 s12260 :: ( 11 Word.word)option)) of Some (imm) => True | _ => False )) then (case (s1798 s12260 :: ( 11 Word.word) option) of (Some (imm)) => return (C_J imm) ) else if ((case ((s1806 s12260 :: (( 3 Word.word * 8 Word.word))option)) of Some ((rs, imm)) => True | _ => False )) then (case (s1806 s12260 :: (( 3 Word.word * 8 Word.word)) option) of (Some ((rs, imm))) => return (C_BEQZ (imm, rs)) ) else if ((case ((s1818 s12260 :: (( 3 Word.word * 8 Word.word))option)) of Some ((rs, imm)) => True | _ => False )) then (case (s1818 s12260 :: (( 3 Word.word * 8 Word.word)) option) of (Some ((rs, imm))) => return (C_BNEZ (imm, rs)) ) else if ((case ((s1830 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rsd, shamt)) => ((((shamt \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))) | _ => False )) then (case (s1830 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rsd, shamt))) => return (C_SLLI (shamt, rsd)) ) else if ((case ((s1842 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, uimm)) => (rd \ zreg) | _ => False )) then (case (s1842 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, uimm))) => return (C_LWSP (uimm, rd)) ) else if ((case ((s1854 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, uimm)) => ((((rd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))) | _ => False )) then (case (s1854 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, uimm))) => return (C_LDSP (uimm, rd)) ) else if ((case ((s1866 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, uimm)) => True | _ => False )) then (case (s1866 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, uimm))) => return (C_SWSP (uimm, rd)) ) else if ((case ((s1878 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rs2, uimm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1878 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rs2, uimm))) => return (C_SDSP (uimm, rs2)) ) else if ((case ((s1890 s12260 :: ( 5 Word.word)option)) of Some (rs1) => (rs1 \ zreg) | _ => False )) then (case (s1890 s12260 :: ( 5 Word.word) option) of (Some (rs1)) => return (C_JR rs1) ) else if ((case ((s1898 s12260 :: ( 5 Word.word)option)) of Some (rs1) => (rs1 \ zreg) | _ => False )) then (case (s1898 s12260 :: ( 5 Word.word) option) of (Some (rs1)) => return (C_JALR rs1) ) else if ((case ((s1906 s12260 :: (( 5 Word.word * 5 Word.word))option)) of Some ((rd, rs2)) => ((((rd \ zreg))) \ (((rs2 \ zreg)))) | _ => False )) then (case (s1906 s12260 :: (( 5 Word.word * 5 Word.word)) option) of (Some ((rd, rs2))) => return (C_MV (rd, rs2)) ) else if (((s12260 = (''c.ebreak'')))) then return (C_EBREAK () ) else if ((case ((s1918 s12260 :: (( 5 Word.word * 5 Word.word))option)) of Some ((rsd, rs2)) => ((((rsd \ zreg))) \ (((rs2 \ zreg)))) | _ => False )) then (case (s1918 s12260 :: (( 5 Word.word * 5 Word.word)) option) of (Some ((rsd, rs2))) => return (C_ADD (rsd, rs2)) ) else if ((case ((s1930 s12260 :: ((bool * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((high, signed1, signed2, rd, rs1, rs2)) => True | _ => False )) then (case (s1930 s12260 :: (( bool * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((high, signed1, signed2, rd, rs1, rs2))) => return (MUL (rs2, rs1, rd, high, signed1, signed2)) ) else if ((case ((s1947 s12260 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((s, rd, rs1, rs2)) => True | _ => False )) then (case (s1947 s12260 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((s, rd, rs1, rs2))) => return (DIV (rs2, rs1, rd, s)) ) else if ((case ((s1965 s12260 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((s, rd, rs1, rs2)) => True | _ => False )) then (case (s1965 s12260 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((s, rd, rs1, rs2))) => return (REM (rs2, rs1, rd, s)) ) else if ((case ((s1983 s12260 :: (( 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((rd, rs1, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1983 s12260 :: (( 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((rd, rs1, rs2))) => return (MULW (rs2, rs1, rd)) ) else if ((case ((s1999 s12260 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((s, rd, rs1, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s1999 s12260 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((s, rd, rs1, rs2))) => return (DIVW (rs2, rs1, rd, s)) ) else if ((case ((s2018 s12260 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((s, rd, rs1, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s2018 s12260 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((s, rd, rs1, rs2))) => return (REMW (rs2, rs1, rd, s)) ) else if ((case ((s2037 s12260 :: ((csrop * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((op1, rd, csr, rs1)) => True | _ => False )) then (case (s2037 s12260 :: (( csrop * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((op1, rd, csr, rs1))) => return (CSR (csr, rs1, rd, True, op1)) ) else if ((case ((s2055 s12260 :: ((csrop * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((op1, rd, csr, rs1)) => True | _ => False )) then (case (s2055 s12260 :: (( csrop * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((op1, rd, csr, rs1))) => return (CSR (csr, rs1, rd, False, op1)) ) else if (((s12260 = (''uret'')))) then return (URET () ) else if ((case ((s2072 s12260 :: ( 6 Word.word)option)) of Some (imm) => True | _ => False )) then (case (s2072 s12260 :: ( 6 Word.word) option) of (Some (imm)) => return (C_NOP_HINT imm) ) else if ((case ((s2078 s12260 :: ( 5 Word.word)option)) of Some (rsd) => (rsd \ zreg) | _ => False )) then (case (s2078 s12260 :: ( 5 Word.word) option) of (Some (rsd)) => return (C_ADDI_HINT rsd) ) else if ((case ((s2084 s12260 :: ( 6 Word.word)option)) of Some (imm) => True | _ => False )) then (case (s2084 s12260 :: ( 6 Word.word) option) of (Some (imm)) => return (C_LI_HINT imm) ) else if ((case ((s2090 s12260 :: ( 6 Word.word)option)) of Some (imm) => (imm \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s2090 s12260 :: ( 6 Word.word) option) of (Some (imm)) => return (C_LUI_HINT imm) ) else if ((case ((s2096 s12260 :: ( 5 Word.word)option)) of Some (rs2) => (rs2 \ zreg) | _ => False )) then (case (s2096 s12260 :: ( 5 Word.word) option) of (Some (rs2)) => return (C_MV_HINT rs2) ) else if ((case ((s2102 s12260 :: ( 5 Word.word)option)) of Some (rs2) => (rs2 \ zreg) | _ => False )) then (case (s2102 s12260 :: ( 5 Word.word) option) of (Some (rs2)) => return (C_ADD_HINT rs2) ) else if ((case ((s2108 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rsd, shamt)) => ((((shamt = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))) | _ => False )) then (case (s2108 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rsd, shamt))) => return (C_SLLI_HINT (shamt, rsd)) ) else if ((case ((s2117 s12260 :: ( 3 Word.word)option)) of Some (rsd) => True | _ => False )) then (case (s2117 s12260 :: ( 3 Word.word) option) of (Some (rsd)) => return (C_SRLI_HINT rsd) ) else if ((case ((s2123 s12260 :: ( 3 Word.word)option)) of Some (rsd) => True | _ => False )) then (case (s2123 s12260 :: ( 3 Word.word) option) of (Some (rsd)) => return (C_SRAI_HINT rsd) ) else if ((case ((s2129 s12260 :: (( 4 Word.word * 4 Word.word * 5 Word.word * 5 Word.word * 4 Word.word))option)) of Some ((pred, succ, rs, rd, fm)) => (((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word)))))))) | _ => False )) then (case (s2129 s12260 :: (( 4 Word.word * 4 Word.word * 5 Word.word * 5 Word.word * 4 Word.word)) option) of (Some ((pred, succ, rs, rd, fm))) => return (FENCE_RESERVED (fm, pred, succ, rs, rd)) ) else if ((case ((s2147 s12260 :: (( 5 Word.word * 5 Word.word * 12 Word.word))option)) of Some ((rd, rs, imm)) => ((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg))))))) | _ => False )) then (case (s2147 s12260 :: (( 5 Word.word * 5 Word.word * 12 Word.word)) option) of (Some ((rd, rs, imm))) => return (FENCEI_RESERVED (imm, rs, rd)) ) else if ((case ((s2159 s12260 :: ((word_width * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((width, rd, imm, rs1)) => True | _ => False )) then (case (s2159 s12260 :: (( word_width * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((width, rd, imm, rs1))) => return (LOAD_FP (imm, rs1, rd, width)) ) else if ((case ((s2183 s12260 :: ((word_width * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((width, rs2, imm, rs1)) => True | _ => False )) then (case (s2183 s12260 :: (( word_width * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((width, rs2, imm, rs1))) => return (STORE_FP (imm, rs2, rs1, width)) ) else if ((case ((s2207 s12260 :: ((f_madd_op_S * 5 Word.word * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((op1, rd, rs1, rs2, rs3, rm)) => True | _ => False )) then (case (s2207 s12260 :: (( f_madd_op_S * 5 Word.word * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((op1, rd, rs1, rs2, rs3, rm))) => return (F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, op1)) ) else if ((case ((s2232 s12260 :: ((f_bin_rm_op_S * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((op1, rd, rs1, rs2, rm)) => True | _ => False )) then (case (s2232 s12260 :: (( f_bin_rm_op_S * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((op1, rd, rs1, rs2, rm))) => return (F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, op1)) ) else if ((case ((s2253 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FSQRT_S, rd, rs1, rm)) => True | _ => False )) then (case (s2253 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FSQRT_S, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FSQRT_S)) ) else if ((case ((s2270 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_W_S, rd, rs1, rm)) => True | _ => False )) then (case (s2270 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_W_S, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_W_S)) ) else if ((case ((s2287 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_WU_S, rd, rs1, rm)) => True | _ => False )) then (case (s2287 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_WU_S, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_WU_S)) ) else if ((case ((s2304 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_S_W, rd, rs1, rm)) => True | _ => False )) then (case (s2304 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_S_W, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_W)) ) else if ((case ((s2321 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_S_WU, rd, rs1, rm)) => True | _ => False )) then (case (s2321 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_S_WU, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_WU)) ) else if ((case ((s2338 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_L_S, rd, rs1, rm)) => True | _ => False )) then (case (s2338 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_L_S, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_L_S)) ) else if ((case ((s2355 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_LU_S, rd, rs1, rm)) => True | _ => False )) then (case (s2355 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_LU_S, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_LU_S)) ) else if ((case ((s2372 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_S_L, rd, rs1, rm)) => True | _ => False )) then (case (s2372 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_S_L, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_L)) ) else if ((case ((s2389 s12260 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_S_LU, rd, rs1, rm)) => True | _ => False )) then (case (s2389 s12260 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_S_LU, rd, rs1, rm))) => return (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_LU)) ) else if ((case ((s2406 s12260 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FSGNJ_S, rd, rs1, rs2)) => True | _ => False )) then (case (s2406 s12260 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FSGNJ_S, rd, rs1, rs2))) => return (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJ_S)) ) else if ((case ((s2423 s12260 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FSGNJN_S, rd, rs1, rs2)) => True | _ => False )) then (case (s2423 s12260 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FSGNJN_S, rd, rs1, rs2))) => return (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJN_S)) ) else if ((case ((s2440 s12260 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FSGNJX_S, rd, rs1, rs2)) => True | _ => False )) then (case (s2440 s12260 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FSGNJX_S, rd, rs1, rs2))) => return (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJX_S)) ) else if ((case ((s2457 s12260 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FMIN_S, rd, rs1, rs2)) => True | _ => False )) then (case (s2457 s12260 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FMIN_S, rd, rs1, rs2))) => return (F_BIN_TYPE_S (rs2, rs1, rd, FMIN_S)) ) else if ((case ((s2474 s12260 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FMAX_S, rd, rs1, rs2)) => True | _ => False )) then (case (s2474 s12260 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FMAX_S, rd, rs1, rs2))) => return (F_BIN_TYPE_S (rs2, rs1, rd, FMAX_S)) ) else if ((case ((s2491 s12260 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FEQ_S, rd, rs1, rs2)) => True | _ => False )) then (case (s2491 s12260 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FEQ_S, rd, rs1, rs2))) => return (F_BIN_TYPE_S (rs2, rs1, rd, FEQ_S)) ) else if ((case ((s2508 s12260 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FLT_S, rd, rs1, rs2)) => True | _ => False )) then (case (s2508 s12260 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FLT_S, rd, rs1, rs2))) => return (F_BIN_TYPE_S (rs2, rs1, rd, FLT_S)) ) else if ((case ((s2525 s12260 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FLE_S, rd, rs1, rs2)) => True | _ => False )) then (case (s2525 s12260 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FLE_S, rd, rs1, rs2))) => return (F_BIN_TYPE_S (rs2, rs1, rd, FLE_S)) ) else if ((case ((s2542 s12260 :: ((f_un_op_S * 5 Word.word * 5 Word.word))option)) of Some ((FMV_X_W, rd, rs1)) => True | _ => False )) then (case (s2542 s12260 :: (( f_un_op_S * 5 Word.word * 5 Word.word)) option) of (Some ((FMV_X_W, rd, rs1))) => return (F_UN_TYPE_S (rs1, rd, FMV_X_W)) ) else if ((case ((s2555 s12260 :: ((f_un_op_S * 5 Word.word * 5 Word.word))option)) of Some ((FMV_W_X, rd, rs1)) => True | _ => False )) then (case (s2555 s12260 :: (( f_un_op_S * 5 Word.word * 5 Word.word)) option) of (Some ((FMV_W_X, rd, rs1))) => return (F_UN_TYPE_S (rs1, rd, FMV_W_X)) ) else if ((case ((s2568 s12260 :: ((f_un_op_S * 5 Word.word * 5 Word.word))option)) of Some ((FCLASS_S, rd, rs1)) => True | _ => False )) then (case (s2568 s12260 :: (( f_un_op_S * 5 Word.word * 5 Word.word)) option) of (Some ((FCLASS_S, rd, rs1))) => return (F_UN_TYPE_S (rs1, rd, FCLASS_S)) ) else if ((case ((s2581 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, imm)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s2581 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, imm))) => return (C_FLWSP (imm, rd)) ) else if ((case ((s2593 s12260 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, uimm)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s2593 s12260 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, uimm))) => return (C_FSWSP (uimm, rd)) ) else if ((case ((s2605 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rdc, rsc, uimm)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s2605 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rdc, rsc, uimm))) => return (C_FLW (uimm, rsc, rdc)) ) else if ((case ((s2621 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rsc1, rsc2, uimm)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s2621 s12260 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rsc1, rsc2, uimm))) => return (C_FSW (uimm, rsc1, rsc2)) ) else if ((case ((s2637 s12260 :: ( 32 Word.word)option)) of Some (s) => True | _ => False )) then (case (s2637 s12260 :: ( 32 Word.word) option) of (Some (s)) => return (ILLEGAL s) ) else if ((case ((s2645 s12260 :: ( 16 Word.word)option)) of Some (s) => True | _ => False )) then (case (s2645 s12260 :: ( 16 Word.word) option) of (Some (s)) => return (C_ILLEGAL s) ) else assert_exp False (''Pattern match failure at unknown location'') \ exit0 () ))\ for arg1 :: " string " definition assembly_forwards_matches :: \ ast \ bool \ where \ assembly_forwards_matches ast = ( (let arg1 = ast in (case arg1 of UTYPE ((imm, rd, op1)) => True | RISCV_JAL ((imm, rd)) => True | RISCV_JALR ((imm, rs1, rd)) => True | BTYPE ((imm, rs2, rs1, op1)) => True | ITYPE ((imm, rs1, rd, op1)) => True | SHIFTIOP ((shamt, rs1, rd, op1)) => True | RTYPE ((rs2, rs1, rd, op1)) => True | LOAD ((imm, rs1, rd, is_unsigned, size1, aq, rl)) => True | STORE ((imm, rs2, rs1, size1, aq, rl)) => True | ADDIW ((imm, rs1, rd)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | SHIFTW ((shamt, rs1, rd, op1)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | RTYPEW ((rs2, rs1, rd, op1)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | SHIFTIWOP ((shamt, rs1, rd, op1)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | FENCE ((pred, succ)) => True | FENCE_TSO ((pred, succ)) => True | FENCEI (_) => True | ECALL (_) => True | MRET (_) => True | SRET (_) => True | EBREAK (_) => True | WFI (_) => True | SFENCE_VMA ((rs1, rs2)) => True | LOADRES ((aq, rl, rs1, size1, rd)) => True | STORECON ((aq, rl, rs2, rs1, size1, rd)) => True | AMO ((op1, aq, rl, rs2, rs1, width, rd)) => True | C_NOP (_) => True | C_ADDI4SPN ((rdc, nzimm)) => if (((nzimm \ ( 0x00 :: 8 Word.word)))) then True else False | C_LW ((uimm, rsc, rdc)) => True | C_LD ((uimm, rsc, rdc)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | C_SW ((uimm, rsc1, rsc2)) => True | C_SD ((uimm, rsc1, rsc2)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | C_ADDI ((nzi, rsd)) => if ((((((nzi \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))))) then True else False | C_JAL (imm) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then True else False | C_ADDIW ((imm, rsd)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | C_LI ((imm, rd)) => if (((rd \ zreg))) then True else False | C_ADDI16SP (imm) => if (((imm \ ( 0b000000 :: 6 Word.word)))) then True else False | C_LUI ((imm, rd)) => if ((((((rd \ zreg))) \ ((((((rd \ sp))) \ (((imm \ ( 0b000000 :: 6 Word.word)))))))))) then True else False | C_SRLI ((shamt, rsd)) => if (((shamt \ ( 0b000000 :: 6 Word.word)))) then True else False | C_SRAI ((shamt, rsd)) => if (((shamt \ ( 0b000000 :: 6 Word.word)))) 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 ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | C_ADDW ((rsd, rs2)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | C_J (imm) => True | C_BEQZ ((imm, rs)) => True | C_BNEZ ((imm, rs)) => True | C_SLLI ((shamt, rsd)) => if ((((((shamt \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))))) then True else False | C_LWSP ((uimm, rd)) => if (((rd \ zreg))) then True else False | C_LDSP ((uimm, rd)) => if ((((((rd \ zreg))) \ ((((( 32 :: int):: ii) = (( 64 :: int):: ii))))))) then True else False | C_SWSP ((uimm, rd)) => True | C_SDSP ((uimm, rs2)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | C_JR (rs1) => if (((rs1 \ zreg))) then True else False | C_JALR (rs1) => if (((rs1 \ zreg))) then True else False | C_MV ((rd, rs2)) => if ((((((rd \ zreg))) \ (((rs2 \ zreg)))))) then True else False | C_EBREAK (_) => True | C_ADD ((rsd, rs2)) => if ((((((rsd \ zreg))) \ (((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 ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | DIVW ((rs2, rs1, rd, s)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | REMW ((rs2, rs1, rd, s)) => if ((((( 32 :: int):: ii) = (( 64 :: int):: ii)))) then True else False | CSR ((csr, rs1, rd, True, op1)) => True | CSR ((csr, rs1, rd, False, op1)) => True | URET (_) => True | C_NOP_HINT (imm) => True | C_ADDI_HINT (rsd) => if (((rsd \ zreg))) then True else False | C_LI_HINT (imm) => True | C_LUI_HINT (imm) => if (((imm \ ( 0b000000 :: 6 Word.word)))) then True else False | C_MV_HINT (rs2) => if (((rs2 \ zreg))) then True else False | C_ADD_HINT (rs2) => if (((rs2 \ zreg))) then True else False | C_SLLI_HINT ((shamt, rsd)) => if ((((((shamt = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))))) then True else False | C_SRLI_HINT (rsd) => True | C_SRAI_HINT (rsd) => True | FENCE_RESERVED ((fm, pred, succ, rs, rd)) => if (((((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word)))))))))) then True else False | FENCEI_RESERVED ((imm, rs, rd)) => if ((((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg))))))))) then True else False | LOAD_FP ((imm, rs1, rd, width)) => True | STORE_FP ((imm, rs2, rs1, width)) => True | F_MADD_TYPE_S ((rs3, rs2, rs1, rm, rd, op1)) => True | F_BIN_RM_TYPE_S ((rs2, rs1, rm, rd, op1)) => 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 ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then True else False | C_FSWSP ((uimm, rd)) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then True else False | C_FLW ((uimm, rsc, rdc)) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then True else False | C_FSW ((uimm, rsc1, rsc2)) => if ((((( 32 :: int):: ii) = (( 32 :: int):: ii)))) then True else False | ILLEGAL (s) => True | C_ILLEGAL (s) => True )))\ for ast :: " ast " \ \\val _s4084_ : string -> maybe (mword ty16)\\ definition s4084 :: \ string \((16)Word.word)option \ where \ s4084 s40850 = ( (let s40860 = s40850 in if ((string_startswith s40860 (''c.illegal''))) then (case ((string_drop s40860 ((string_length (''c.illegal''))))) of s40870 => (case ((spc_matches_prefix0 s40870)) of Some ((_, s40880)) => (case ((string_drop s40870 s40880)) of s40890 => (case ((hex_bits_16_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s40890 :: (( 16 Word.word * ii)) option)) of Some ((s, s40900)) => (let p00 = (string_drop s40890 s40900) in if (((p00 = ('''')))) then Some s else None) | _ => None ) ) | _ => None ) ) else None))\ for s40850 :: " string " \ \\val _s4076_ : string -> maybe (mword ty32)\\ definition s4076 :: \ string \((32)Word.word)option \ where \ s4076 s40770 = ( (let s40780 = s40770 in if ((string_startswith s40780 (''illegal''))) then (case ((string_drop s40780 ((string_length (''illegal''))))) of s40790 => (case ((spc_matches_prefix0 s40790)) of Some ((_, s40800)) => (case ((string_drop s40790 s40800)) of s40810 => (case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s40810 :: (( 32 Word.word * ii)) option)) of Some ((s, s40820)) => (let p00 = (string_drop s40810 s40820) in if (((p00 = ('''')))) then Some s else None) | _ => None ) ) | _ => None ) ) else None))\ for s40770 :: " string " \ \\val _s4060_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s4060 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s4060 s40610 = ( (let s40620 = s40610 in if ((string_startswith s40620 (''c.fsw''))) then (case ((string_drop s40620 ((string_length (''c.fsw''))))) of s40630 => (case ((spc_matches_prefix0 s40630)) of Some ((_, s40640)) => (case ((string_drop s40630 s40640)) of s40650 => (case ((creg_name_matches_prefix s40650 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s40660)) => (case ((string_drop s40650 s40660)) of s40670 => (case ((sep_matches_prefix s40670)) of Some ((_, s40680)) => (case ((string_drop s40670 s40680)) of s40690 => (case ((creg_name_matches_prefix s40690 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s40700)) => (case ((string_drop s40690 s40700)) of s40710 => (case ((sep_matches_prefix s40710)) of Some ((_, s40720)) => (case ((string_drop s40710 s40720)) of s40730 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s40730 :: (( 7 Word.word * ii)) option)) of Some ((v__1196, s40740)) => if (((((subrange_vec_dec v__1196 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1196 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1196 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s40730 s40740) in if (((p00 = ('''')))) then Some (rsc1, rsc2, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s40610 :: " string " \ \\val _s4044_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s4044 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s4044 s40450 = ( (let s40460 = s40450 in if ((string_startswith s40460 (''c.flw''))) then (case ((string_drop s40460 ((string_length (''c.flw''))))) of s40470 => (case ((spc_matches_prefix0 s40470)) of Some ((_, s40480)) => (case ((string_drop s40470 s40480)) of s40490 => (case ((creg_name_matches_prefix s40490 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s40500)) => (case ((string_drop s40490 s40500)) of s40510 => (case ((sep_matches_prefix s40510)) of Some ((_, s40520)) => (case ((string_drop s40510 s40520)) of s40530 => (case ((creg_name_matches_prefix s40530 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s40540)) => (case ((string_drop s40530 s40540)) of s40550 => (case ((sep_matches_prefix s40550)) of Some ((_, s40560)) => (case ((string_drop s40550 s40560)) of s40570 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s40570 :: (( 7 Word.word * ii)) option)) of Some ((v__1198, s40580)) => if (((((subrange_vec_dec v__1198 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1198 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1198 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s40570 s40580) in if (((p00 = ('''')))) then Some (rdc, rsc, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s40450 :: " string " \ \\val _s4032_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s4032 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s4032 s40330 = ( (let s40340 = s40330 in if ((string_startswith s40340 (''c.fswsp''))) then (case ((string_drop s40340 ((string_length (''c.fswsp''))))) of s40350 => (case ((spc_matches_prefix0 s40350)) of Some ((_, s40360)) => (case ((string_drop s40350 s40360)) of s40370 => (case ((reg_name_matches_prefix s40370 :: (( 5 Word.word * ii)) option)) of Some ((rd, s40380)) => (case ((string_drop s40370 s40380)) of s40390 => (case ((sep_matches_prefix s40390)) of Some ((_, s40400)) => (case ((string_drop s40390 s40400)) of s40410 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s40410 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s40420)) => (let p00 = (string_drop s40410 s40420) in if (((p00 = ('''')))) then Some (rd, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s40330 :: " string " \ \\val _s4020_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s4020 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s4020 s40210 = ( (let s40220 = s40210 in if ((string_startswith s40220 (''c.flwsp''))) then (case ((string_drop s40220 ((string_length (''c.flwsp''))))) of s40230 => (case ((spc_matches_prefix0 s40230)) of Some ((_, s40240)) => (case ((string_drop s40230 s40240)) of s40250 => (case ((reg_name_matches_prefix s40250 :: (( 5 Word.word * ii)) option)) of Some ((rd, s40260)) => (case ((string_drop s40250 s40260)) of s40270 => (case ((sep_matches_prefix s40270)) of Some ((_, s40280)) => (case ((string_drop s40270 s40280)) of s40290 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s40290 :: (( 6 Word.word * ii)) option)) of Some ((imm, s40300)) => (let p00 = (string_drop s40290 s40300) in if (((p00 = ('''')))) then Some (rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s40210 :: " string " \ \\val _s4007_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5))\\ definition s4007 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word)option \ where \ s4007 s40090 = ( (case ((f_un_type_mnemonic_S_matches_prefix s40090)) of Some ((FCLASS_S, s40100)) => (case ((string_drop s40090 s40100)) of s40110 => (case ((spc_matches_prefix0 s40110)) of Some ((_, s40120)) => (case ((string_drop s40110 s40120)) of s40130 => (case ((reg_name_matches_prefix s40130 :: (( 5 Word.word * ii)) option)) of Some ((rd, s40140)) => (case ((string_drop s40130 s40140)) of s40150 => (case ((sep_matches_prefix s40150)) of Some ((_, s40160)) => (case ((string_drop s40150 s40160)) of s40170 => (case ((freg_name_matches_prefix s40170 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s40180)) => (let p00 = (string_drop s40170 s40180) in if (((p00 = ('''')))) then Some (FCLASS_S, rd, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s40090 :: " string " \ \\val _s3994_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5))\\ definition s3994 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word)option \ where \ s3994 s39960 = ( (case ((f_un_type_mnemonic_S_matches_prefix s39960)) of Some ((FMV_W_X, s39970)) => (case ((string_drop s39960 s39970)) of s39980 => (case ((spc_matches_prefix0 s39980)) of Some ((_, s39990)) => (case ((string_drop s39980 s39990)) of s40000 => (case ((freg_name_matches_prefix s40000 :: (( 5 Word.word * ii)) option)) of Some ((rd, s40010)) => (case ((string_drop s40000 s40010)) of s40020 => (case ((sep_matches_prefix s40020)) of Some ((_, s40030)) => (case ((string_drop s40020 s40030)) of s40040 => (case ((reg_name_matches_prefix s40040 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s40050)) => (let p00 = (string_drop s40040 s40050) in if (((p00 = ('''')))) then Some (FMV_W_X, rd, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s39960 :: " string " \ \\val _s3981_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5))\\ definition s3981 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word)option \ where \ s3981 s39830 = ( (case ((f_un_type_mnemonic_S_matches_prefix s39830)) of Some ((FMV_X_W, s39840)) => (case ((string_drop s39830 s39840)) of s39850 => (case ((spc_matches_prefix0 s39850)) of Some ((_, s39860)) => (case ((string_drop s39850 s39860)) of s39870 => (case ((reg_name_matches_prefix s39870 :: (( 5 Word.word * ii)) option)) of Some ((rd, s39880)) => (case ((string_drop s39870 s39880)) of s39890 => (case ((sep_matches_prefix s39890)) of Some ((_, s39900)) => (case ((string_drop s39890 s39900)) of s39910 => (case ((freg_name_matches_prefix s39910 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s39920)) => (let p00 = (string_drop s39910 s39920) in if (((p00 = ('''')))) then Some (FMV_X_W, rd, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s39830 :: " string " \ \\val _s3964_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s3964 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3964 s39660 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s39660)) of Some ((FLE_S, s39670)) => (case ((string_drop s39660 s39670)) of s39680 => (case ((spc_matches_prefix0 s39680)) of Some ((_, s39690)) => (case ((string_drop s39680 s39690)) of s39700 => (case ((reg_name_matches_prefix s39700 :: (( 5 Word.word * ii)) option)) of Some ((rd, s39710)) => (case ((string_drop s39700 s39710)) of s39720 => (case ((sep_matches_prefix s39720)) of Some ((_, s39730)) => (case ((string_drop s39720 s39730)) of s39740 => (case ((freg_name_matches_prefix s39740 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s39750)) => (case ((string_drop s39740 s39750)) of s39760 => (case ((sep_matches_prefix s39760)) of Some ((_, s39770)) => (case ((string_drop s39760 s39770)) of s39780 => (case ((freg_name_matches_prefix s39780 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s39790)) => (let p00 = (string_drop s39780 s39790) in if (((p00 = ('''')))) then Some (FLE_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s39660 :: " string " \ \\val _s3947_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s3947 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3947 s39490 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s39490)) of Some ((FLT_S, s39500)) => (case ((string_drop s39490 s39500)) of s39510 => (case ((spc_matches_prefix0 s39510)) of Some ((_, s39520)) => (case ((string_drop s39510 s39520)) of s39530 => (case ((reg_name_matches_prefix s39530 :: (( 5 Word.word * ii)) option)) of Some ((rd, s39540)) => (case ((string_drop s39530 s39540)) of s39550 => (case ((sep_matches_prefix s39550)) of Some ((_, s39560)) => (case ((string_drop s39550 s39560)) of s39570 => (case ((freg_name_matches_prefix s39570 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s39580)) => (case ((string_drop s39570 s39580)) of s39590 => (case ((sep_matches_prefix s39590)) of Some ((_, s39600)) => (case ((string_drop s39590 s39600)) of s39610 => (case ((freg_name_matches_prefix s39610 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s39620)) => (let p00 = (string_drop s39610 s39620) in if (((p00 = ('''')))) then Some (FLT_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s39490 :: " string " \ \\val _s3930_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s3930 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3930 s39320 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s39320)) of Some ((FEQ_S, s39330)) => (case ((string_drop s39320 s39330)) of s39340 => (case ((spc_matches_prefix0 s39340)) of Some ((_, s39350)) => (case ((string_drop s39340 s39350)) of s39360 => (case ((reg_name_matches_prefix s39360 :: (( 5 Word.word * ii)) option)) of Some ((rd, s39370)) => (case ((string_drop s39360 s39370)) of s39380 => (case ((sep_matches_prefix s39380)) of Some ((_, s39390)) => (case ((string_drop s39380 s39390)) of s39400 => (case ((freg_name_matches_prefix s39400 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s39410)) => (case ((string_drop s39400 s39410)) of s39420 => (case ((sep_matches_prefix s39420)) of Some ((_, s39430)) => (case ((string_drop s39420 s39430)) of s39440 => (case ((freg_name_matches_prefix s39440 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s39450)) => (let p00 = (string_drop s39440 s39450) in if (((p00 = ('''')))) then Some (FEQ_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s39320 :: " string " \ \\val _s3913_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s3913 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3913 s39150 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s39150)) of Some ((FMAX_S, s39160)) => (case ((string_drop s39150 s39160)) of s39170 => (case ((spc_matches_prefix0 s39170)) of Some ((_, s39180)) => (case ((string_drop s39170 s39180)) of s39190 => (case ((freg_name_matches_prefix s39190 :: (( 5 Word.word * ii)) option)) of Some ((rd, s39200)) => (case ((string_drop s39190 s39200)) of s39210 => (case ((sep_matches_prefix s39210)) of Some ((_, s39220)) => (case ((string_drop s39210 s39220)) of s39230 => (case ((freg_name_matches_prefix s39230 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s39240)) => (case ((string_drop s39230 s39240)) of s39250 => (case ((sep_matches_prefix s39250)) of Some ((_, s39260)) => (case ((string_drop s39250 s39260)) of s39270 => (case ((freg_name_matches_prefix s39270 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s39280)) => (let p00 = (string_drop s39270 s39280) in if (((p00 = ('''')))) then Some (FMAX_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s39150 :: " string " \ \\val _s3896_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s3896 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3896 s38980 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s38980)) of Some ((FMIN_S, s38990)) => (case ((string_drop s38980 s38990)) of s39000 => (case ((spc_matches_prefix0 s39000)) of Some ((_, s39010)) => (case ((string_drop s39000 s39010)) of s39020 => (case ((freg_name_matches_prefix s39020 :: (( 5 Word.word * ii)) option)) of Some ((rd, s39030)) => (case ((string_drop s39020 s39030)) of s39040 => (case ((sep_matches_prefix s39040)) of Some ((_, s39050)) => (case ((string_drop s39040 s39050)) of s39060 => (case ((freg_name_matches_prefix s39060 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s39070)) => (case ((string_drop s39060 s39070)) of s39080 => (case ((sep_matches_prefix s39080)) of Some ((_, s39090)) => (case ((string_drop s39080 s39090)) of s39100 => (case ((freg_name_matches_prefix s39100 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s39110)) => (let p00 = (string_drop s39100 s39110) in if (((p00 = ('''')))) then Some (FMIN_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s38980 :: " string " \ \\val _s3879_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s3879 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3879 s38810 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s38810)) of Some ((FSGNJX_S, s38820)) => (case ((string_drop s38810 s38820)) of s38830 => (case ((spc_matches_prefix0 s38830)) of Some ((_, s38840)) => (case ((string_drop s38830 s38840)) of s38850 => (case ((freg_name_matches_prefix s38850 :: (( 5 Word.word * ii)) option)) of Some ((rd, s38860)) => (case ((string_drop s38850 s38860)) of s38870 => (case ((sep_matches_prefix s38870)) of Some ((_, s38880)) => (case ((string_drop s38870 s38880)) of s38890 => (case ((freg_name_matches_prefix s38890 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s38900)) => (case ((string_drop s38890 s38900)) of s38910 => (case ((sep_matches_prefix s38910)) of Some ((_, s38920)) => (case ((string_drop s38910 s38920)) of s38930 => (case ((freg_name_matches_prefix s38930 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s38940)) => (let p00 = (string_drop s38930 s38940) in if (((p00 = ('''')))) then Some (FSGNJX_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s38810 :: " string " \ \\val _s3862_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s3862 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3862 s38640 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s38640)) of Some ((FSGNJN_S, s38650)) => (case ((string_drop s38640 s38650)) of s38660 => (case ((spc_matches_prefix0 s38660)) of Some ((_, s38670)) => (case ((string_drop s38660 s38670)) of s38680 => (case ((freg_name_matches_prefix s38680 :: (( 5 Word.word * ii)) option)) of Some ((rd, s38690)) => (case ((string_drop s38680 s38690)) of s38700 => (case ((sep_matches_prefix s38700)) of Some ((_, s38710)) => (case ((string_drop s38700 s38710)) of s38720 => (case ((freg_name_matches_prefix s38720 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s38730)) => (case ((string_drop s38720 s38730)) of s38740 => (case ((sep_matches_prefix s38740)) of Some ((_, s38750)) => (case ((string_drop s38740 s38750)) of s38760 => (case ((freg_name_matches_prefix s38760 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s38770)) => (let p00 = (string_drop s38760 s38770) in if (((p00 = ('''')))) then Some (FSGNJN_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s38640 :: " string " \ \\val _s3845_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5))\\ definition s3845 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3845 s38470 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s38470)) of Some ((FSGNJ_S, s38480)) => (case ((string_drop s38470 s38480)) of s38490 => (case ((spc_matches_prefix0 s38490)) of Some ((_, s38500)) => (case ((string_drop s38490 s38500)) of s38510 => (case ((freg_name_matches_prefix s38510 :: (( 5 Word.word * ii)) option)) of Some ((rd, s38520)) => (case ((string_drop s38510 s38520)) of s38530 => (case ((sep_matches_prefix s38530)) of Some ((_, s38540)) => (case ((string_drop s38530 s38540)) of s38550 => (case ((freg_name_matches_prefix s38550 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s38560)) => (case ((string_drop s38550 s38560)) of s38570 => (case ((sep_matches_prefix s38570)) of Some ((_, s38580)) => (case ((string_drop s38570 s38580)) of s38590 => (case ((freg_name_matches_prefix s38590 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s38600)) => (let p00 = (string_drop s38590 s38600) in if (((p00 = ('''')))) then Some (FSGNJ_S, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s38470 :: " string " \ \\val _s3828_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3828 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3828 s38300 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s38300)) of Some ((FCVT_S_LU, s38310)) => (case ((string_drop s38300 s38310)) of s38320 => (case ((spc_matches_prefix0 s38320)) of Some ((_, s38330)) => (case ((string_drop s38320 s38330)) of s38340 => (case ((freg_name_matches_prefix s38340 :: (( 5 Word.word * ii)) option)) of Some ((rd, s38350)) => (case ((string_drop s38340 s38350)) of s38360 => (case ((sep_matches_prefix s38360)) of Some ((_, s38370)) => (case ((string_drop s38360 s38370)) of s38380 => (case ((reg_name_matches_prefix s38380 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s38390)) => (case ((string_drop s38380 s38390)) of s38400 => (case ((sep_matches_prefix s38400)) of Some ((_, s38410)) => (case ((string_drop s38400 s38410)) of s38420 => (case ((frm_mnemonic_matches_prefix s38420)) of Some ((rm, s38430)) => (let p00 = (string_drop s38420 s38430) in if (((p00 = ('''')))) then Some (FCVT_S_LU, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s38300 :: " string " \ \\val _s3811_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3811 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3811 s38130 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s38130)) of Some ((FCVT_S_L, s38140)) => (case ((string_drop s38130 s38140)) of s38150 => (case ((spc_matches_prefix0 s38150)) of Some ((_, s38160)) => (case ((string_drop s38150 s38160)) of s38170 => (case ((freg_name_matches_prefix s38170 :: (( 5 Word.word * ii)) option)) of Some ((rd, s38180)) => (case ((string_drop s38170 s38180)) of s38190 => (case ((sep_matches_prefix s38190)) of Some ((_, s38200)) => (case ((string_drop s38190 s38200)) of s38210 => (case ((reg_name_matches_prefix s38210 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s38220)) => (case ((string_drop s38210 s38220)) of s38230 => (case ((sep_matches_prefix s38230)) of Some ((_, s38240)) => (case ((string_drop s38230 s38240)) of s38250 => (case ((frm_mnemonic_matches_prefix s38250)) of Some ((rm, s38260)) => (let p00 = (string_drop s38250 s38260) in if (((p00 = ('''')))) then Some (FCVT_S_L, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s38130 :: " string " \ \\val _s3794_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3794 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3794 s37960 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s37960)) of Some ((FCVT_LU_S, s37970)) => (case ((string_drop s37960 s37970)) of s37980 => (case ((spc_matches_prefix0 s37980)) of Some ((_, s37990)) => (case ((string_drop s37980 s37990)) of s38000 => (case ((reg_name_matches_prefix s38000 :: (( 5 Word.word * ii)) option)) of Some ((rd, s38010)) => (case ((string_drop s38000 s38010)) of s38020 => (case ((sep_matches_prefix s38020)) of Some ((_, s38030)) => (case ((string_drop s38020 s38030)) of s38040 => (case ((freg_name_matches_prefix s38040 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s38050)) => (case ((string_drop s38040 s38050)) of s38060 => (case ((sep_matches_prefix s38060)) of Some ((_, s38070)) => (case ((string_drop s38060 s38070)) of s38080 => (case ((frm_mnemonic_matches_prefix s38080)) of Some ((rm, s38090)) => (let p00 = (string_drop s38080 s38090) in if (((p00 = ('''')))) then Some (FCVT_LU_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s37960 :: " string " \ \\val _s3777_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3777 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3777 s37790 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s37790)) of Some ((FCVT_L_S, s37800)) => (case ((string_drop s37790 s37800)) of s37810 => (case ((spc_matches_prefix0 s37810)) of Some ((_, s37820)) => (case ((string_drop s37810 s37820)) of s37830 => (case ((reg_name_matches_prefix s37830 :: (( 5 Word.word * ii)) option)) of Some ((rd, s37840)) => (case ((string_drop s37830 s37840)) of s37850 => (case ((sep_matches_prefix s37850)) of Some ((_, s37860)) => (case ((string_drop s37850 s37860)) of s37870 => (case ((freg_name_matches_prefix s37870 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s37880)) => (case ((string_drop s37870 s37880)) of s37890 => (case ((sep_matches_prefix s37890)) of Some ((_, s37900)) => (case ((string_drop s37890 s37900)) of s37910 => (case ((frm_mnemonic_matches_prefix s37910)) of Some ((rm, s37920)) => (let p00 = (string_drop s37910 s37920) in if (((p00 = ('''')))) then Some (FCVT_L_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s37790 :: " string " \ \\val _s3760_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3760 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3760 s37620 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s37620)) of Some ((FCVT_S_WU, s37630)) => (case ((string_drop s37620 s37630)) of s37640 => (case ((spc_matches_prefix0 s37640)) of Some ((_, s37650)) => (case ((string_drop s37640 s37650)) of s37660 => (case ((freg_name_matches_prefix s37660 :: (( 5 Word.word * ii)) option)) of Some ((rd, s37670)) => (case ((string_drop s37660 s37670)) of s37680 => (case ((sep_matches_prefix s37680)) of Some ((_, s37690)) => (case ((string_drop s37680 s37690)) of s37700 => (case ((reg_name_matches_prefix s37700 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s37710)) => (case ((string_drop s37700 s37710)) of s37720 => (case ((sep_matches_prefix s37720)) of Some ((_, s37730)) => (case ((string_drop s37720 s37730)) of s37740 => (case ((frm_mnemonic_matches_prefix s37740)) of Some ((rm, s37750)) => (let p00 = (string_drop s37740 s37750) in if (((p00 = ('''')))) then Some (FCVT_S_WU, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s37620 :: " string " \ \\val _s3743_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3743 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3743 s37450 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s37450)) of Some ((FCVT_S_W, s37460)) => (case ((string_drop s37450 s37460)) of s37470 => (case ((spc_matches_prefix0 s37470)) of Some ((_, s37480)) => (case ((string_drop s37470 s37480)) of s37490 => (case ((freg_name_matches_prefix s37490 :: (( 5 Word.word * ii)) option)) of Some ((rd, s37500)) => (case ((string_drop s37490 s37500)) of s37510 => (case ((sep_matches_prefix s37510)) of Some ((_, s37520)) => (case ((string_drop s37510 s37520)) of s37530 => (case ((reg_name_matches_prefix s37530 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s37540)) => (case ((string_drop s37530 s37540)) of s37550 => (case ((sep_matches_prefix s37550)) of Some ((_, s37560)) => (case ((string_drop s37550 s37560)) of s37570 => (case ((frm_mnemonic_matches_prefix s37570)) of Some ((rm, s37580)) => (let p00 = (string_drop s37570 s37580) in if (((p00 = ('''')))) then Some (FCVT_S_W, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s37450 :: " string " \ \\val _s3726_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3726 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3726 s37280 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s37280)) of Some ((FCVT_WU_S, s37290)) => (case ((string_drop s37280 s37290)) of s37300 => (case ((spc_matches_prefix0 s37300)) of Some ((_, s37310)) => (case ((string_drop s37300 s37310)) of s37320 => (case ((reg_name_matches_prefix s37320 :: (( 5 Word.word * ii)) option)) of Some ((rd, s37330)) => (case ((string_drop s37320 s37330)) of s37340 => (case ((sep_matches_prefix s37340)) of Some ((_, s37350)) => (case ((string_drop s37340 s37350)) of s37360 => (case ((freg_name_matches_prefix s37360 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s37370)) => (case ((string_drop s37360 s37370)) of s37380 => (case ((sep_matches_prefix s37380)) of Some ((_, s37390)) => (case ((string_drop s37380 s37390)) of s37400 => (case ((frm_mnemonic_matches_prefix s37400)) of Some ((rm, s37410)) => (let p00 = (string_drop s37400 s37410) in if (((p00 = ('''')))) then Some (FCVT_WU_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s37280 :: " string " \ \\val _s3709_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3709 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3709 s37110 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s37110)) of Some ((FCVT_W_S, s37120)) => (case ((string_drop s37110 s37120)) of s37130 => (case ((spc_matches_prefix0 s37130)) of Some ((_, s37140)) => (case ((string_drop s37130 s37140)) of s37150 => (case ((reg_name_matches_prefix s37150 :: (( 5 Word.word * ii)) option)) of Some ((rd, s37160)) => (case ((string_drop s37150 s37160)) of s37170 => (case ((sep_matches_prefix s37170)) of Some ((_, s37180)) => (case ((string_drop s37170 s37180)) of s37190 => (case ((freg_name_matches_prefix s37190 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s37200)) => (case ((string_drop s37190 s37200)) of s37210 => (case ((sep_matches_prefix s37210)) of Some ((_, s37220)) => (case ((string_drop s37210 s37220)) of s37230 => (case ((frm_mnemonic_matches_prefix s37230)) of Some ((rm, s37240)) => (let p00 = (string_drop s37230 s37240) in if (((p00 = ('''')))) then Some (FCVT_W_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s37110 :: " string " \ \\val _s3692_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode))\\ definition s3692 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3692 s36940 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s36940)) of Some ((FSQRT_S, s36950)) => (case ((string_drop s36940 s36950)) of s36960 => (case ((spc_matches_prefix0 s36960)) of Some ((_, s36970)) => (case ((string_drop s36960 s36970)) of s36980 => (case ((freg_name_matches_prefix s36980 :: (( 5 Word.word * ii)) option)) of Some ((rd, s36990)) => (case ((string_drop s36980 s36990)) of s37000 => (case ((sep_matches_prefix s37000)) of Some ((_, s37010)) => (case ((string_drop s37000 s37010)) of s37020 => (case ((freg_name_matches_prefix s37020 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s37030)) => (case ((string_drop s37020 s37030)) of s37040 => (case ((sep_matches_prefix s37040)) of Some ((_, s37050)) => (case ((string_drop s37040 s37050)) of s37060 => (case ((frm_mnemonic_matches_prefix s37060)) of Some ((rm, s37070)) => (let p00 = (string_drop s37060 s37070) in if (((p00 = ('''')))) then Some (FSQRT_S, rd, rs1, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s36940 :: " string " \ \\val _s3671_ : string -> maybe ((f_bin_rm_op_S * mword ty5 * mword ty5 * mword ty5 * rounding_mode))\\ definition s3671 :: \ string \(f_bin_rm_op_S*(5)Word.word*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3671 s36730 = ( (case ((f_bin_rm_type_mnemonic_S_matches_prefix s36730)) of Some ((op1, s36740)) => (case ((string_drop s36730 s36740)) of s36750 => (case ((spc_matches_prefix0 s36750)) of Some ((_, s36760)) => (case ((string_drop s36750 s36760)) of s36770 => (case ((freg_name_matches_prefix s36770 :: (( 5 Word.word * ii)) option)) of Some ((rd, s36780)) => (case ((string_drop s36770 s36780)) of s36790 => (case ((sep_matches_prefix s36790)) of Some ((_, s36800)) => (case ((string_drop s36790 s36800)) of s36810 => (case ((freg_name_matches_prefix s36810 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s36820)) => (case ((string_drop s36810 s36820)) of s36830 => (case ((sep_matches_prefix s36830)) of Some ((_, s36840)) => (case ((string_drop s36830 s36840)) of s36850 => (case ((freg_name_matches_prefix s36850 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s36860)) => (case ((string_drop s36850 s36860)) of s36870 => (case ((sep_matches_prefix s36870)) of Some ((_, s36880)) => (case ((string_drop s36870 s36880)) of s36890 => (case ((frm_mnemonic_matches_prefix s36890)) of Some ((rm, s36900)) => (let p00 = (string_drop s36890 s36900) in if (((p00 = ('''')))) then Some (op1, rd, rs1, rs2, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s36730 :: " string " \ \\val _s3646_ : string -> maybe ((f_madd_op_S * mword ty5 * mword ty5 * mword ty5 * mword ty5 * rounding_mode))\\ definition s3646 :: \ string \(f_madd_op_S*(5)Word.word*(5)Word.word*(5)Word.word*(5)Word.word*rounding_mode)option \ where \ s3646 s36480 = ( (case ((f_madd_type_mnemonic_S_matches_prefix s36480)) of Some ((op1, s36490)) => (case ((string_drop s36480 s36490)) of s36500 => (case ((spc_matches_prefix0 s36500)) of Some ((_, s36510)) => (case ((string_drop s36500 s36510)) of s36520 => (case ((freg_name_matches_prefix s36520 :: (( 5 Word.word * ii)) option)) of Some ((rd, s36530)) => (case ((string_drop s36520 s36530)) of s36540 => (case ((sep_matches_prefix s36540)) of Some ((_, s36550)) => (case ((string_drop s36540 s36550)) of s36560 => (case ((freg_name_matches_prefix s36560 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s36570)) => (case ((string_drop s36560 s36570)) of s36580 => (case ((sep_matches_prefix s36580)) of Some ((_, s36590)) => (case ((string_drop s36580 s36590)) of s36600 => (case ((freg_name_matches_prefix s36600 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s36610)) => (case ((string_drop s36600 s36610)) of s36620 => (case ((sep_matches_prefix s36620)) of Some ((_, s36630)) => (case ((string_drop s36620 s36630)) of s36640 => (case ((freg_name_matches_prefix s36640 :: (( 5 Word.word * ii)) option)) of Some ((rs3, s36650)) => (case ((string_drop s36640 s36650)) of s36660 => (case ((sep_matches_prefix s36660)) of Some ((_, s36670)) => (case ((string_drop s36660 s36670)) of s36680 => (case ((frm_mnemonic_matches_prefix s36680)) of Some ((rm, s36690)) => (let p00 = (string_drop s36680 s36690) in if (((p00 = ('''')))) then Some (op1, rd, rs1, rs2, rs3, rm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s36480 :: " string " \ \\val _s3622_ : string -> maybe ((word_width * mword ty5 * mword ty12 * mword ty5))\\ definition s3622 :: \ string \(word_width*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s3622 s36230 = ( (let s36240 = s36230 in if ((string_startswith s36240 (''fs''))) then (case ((string_drop s36240 ((string_length (''fs''))))) of s36250 => (case ((size_mnemonic_matches_prefix s36250)) of Some ((width, s36260)) => (case ((string_drop s36250 s36260)) of s36270 => (case ((spc_matches_prefix0 s36270)) of Some ((_, s36280)) => (case ((string_drop s36270 s36280)) of s36290 => (case ((freg_name_matches_prefix s36290 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s36300)) => (case ((string_drop s36290 s36300)) of s36310 => (case ((sep_matches_prefix s36310)) of Some ((_, s36320)) => (case ((string_drop s36310 s36320)) of s36330 => (case ((hex_bits_12_matches_prefix0 s36330 :: (( 12 Word.word * ii)) option)) of Some ((imm, s36340)) => (case ((string_drop s36330 s36340)) of s36350 => (case ((opt_spc_matches_prefix0 s36350)) of Some ((_, s36360)) => (let s36370 = (string_drop s36350 s36360) in if ((string_startswith s36370 (''(''))) then (case ((string_drop s36370 ((string_length (''(''))))) of s36380 => (case ((opt_spc_matches_prefix0 s36380)) of Some ((_, s36390)) => (case ((string_drop s36380 s36390)) of s36400 => (case ((reg_name_matches_prefix s36400 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s36410)) => (case ((string_drop s36400 s36410)) of s36420 => (case ((opt_spc_matches_prefix0 s36420)) of Some ((_, s36430)) => (let s36440 = (string_drop s36420 s36430) in if ((string_startswith s36440 ('')''))) then (let p00 = (string_drop s36440 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (width, rs2, imm, rs1) else None) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s36230 :: " string " \ \\val _s3598_ : string -> maybe ((word_width * mword ty5 * mword ty12 * mword ty5))\\ definition s3598 :: \ string \(word_width*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s3598 s35990 = ( (let s36000 = s35990 in if ((string_startswith s36000 (''fl''))) then (case ((string_drop s36000 ((string_length (''fl''))))) of s36010 => (case ((size_mnemonic_matches_prefix s36010)) of Some ((width, s36020)) => (case ((string_drop s36010 s36020)) of s36030 => (case ((spc_matches_prefix0 s36030)) of Some ((_, s36040)) => (case ((string_drop s36030 s36040)) of s36050 => (case ((freg_name_matches_prefix s36050 :: (( 5 Word.word * ii)) option)) of Some ((rd, s36060)) => (case ((string_drop s36050 s36060)) of s36070 => (case ((sep_matches_prefix s36070)) of Some ((_, s36080)) => (case ((string_drop s36070 s36080)) of s36090 => (case ((hex_bits_12_matches_prefix0 s36090 :: (( 12 Word.word * ii)) option)) of Some ((imm, s36100)) => (case ((string_drop s36090 s36100)) of s36110 => (case ((opt_spc_matches_prefix0 s36110)) of Some ((_, s36120)) => (let s36130 = (string_drop s36110 s36120) in if ((string_startswith s36130 (''(''))) then (case ((string_drop s36130 ((string_length (''(''))))) of s36140 => (case ((opt_spc_matches_prefix0 s36140)) of Some ((_, s36150)) => (case ((string_drop s36140 s36150)) of s36160 => (case ((reg_name_matches_prefix s36160 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s36170)) => (case ((string_drop s36160 s36170)) of s36180 => (case ((opt_spc_matches_prefix0 s36180)) of Some ((_, s36190)) => (let s36200 = (string_drop s36180 s36190) in if ((string_startswith s36200 ('')''))) then (let p00 = (string_drop s36200 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (width, rd, imm, rs1) else None) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s35990 :: " string " \ \\val _s3586_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12))\\ definition s3586 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word)option \ where \ s3586 s35870 = ( (let s35880 = s35870 in if ((string_startswith s35880 (''fence.i.reserved.''))) then (case ((string_drop s35880 ((string_length (''fence.i.reserved.''))))) of s35890 => (case ((reg_name_matches_prefix s35890 :: (( 5 Word.word * ii)) option)) of Some ((rd, s35900)) => (let s35910 = (string_drop s35890 s35900) in if ((string_startswith s35910 (''.''))) then (case ((string_drop s35910 ((string_length (''.''))))) of s35920 => (case ((reg_name_matches_prefix s35920 :: (( 5 Word.word * ii)) option)) of Some ((rs, s35930)) => (let s35940 = (string_drop s35920 s35930) in if ((string_startswith s35940 (''.''))) then (case ((string_drop s35940 ((string_length (''.''))))) of s35950 => (case ((hex_bits_12_matches_prefix0 s35950 :: (( 12 Word.word * ii)) option)) of Some ((imm, s35960)) => (let p00 = (string_drop s35950 s35960) in if (((p00 = ('''')))) then Some (rd, rs, imm) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None))\ for s35870 :: " string " \ \\val _s3568_ : string -> maybe ((mword ty4 * mword ty4 * mword ty5 * mword ty5 * mword ty4))\\ definition s3568 :: \ string \((4)Word.word*(4)Word.word*(5)Word.word*(5)Word.word*(4)Word.word)option \ where \ s3568 s35690 = ( (let s35700 = s35690 in if ((string_startswith s35700 (''fence.reserved.''))) then (case ((string_drop s35700 ((string_length (''fence.reserved.''))))) of s35710 => (case ((fence_bits_matches_prefix s35710 :: (( 4 Word.word * ii)) option)) of Some ((pred, s35720)) => (let s35730 = (string_drop s35710 s35720) in if ((string_startswith s35730 (''.''))) then (case ((string_drop s35730 ((string_length (''.''))))) of s35740 => (case ((fence_bits_matches_prefix s35740 :: (( 4 Word.word * ii)) option)) of Some ((succ, s35750)) => (let s35760 = (string_drop s35740 s35750) in if ((string_startswith s35760 (''.''))) then (case ((string_drop s35760 ((string_length (''.''))))) of s35770 => (case ((reg_name_matches_prefix s35770 :: (( 5 Word.word * ii)) option)) of Some ((rs, s35780)) => (let s35790 = (string_drop s35770 s35780) in if ((string_startswith s35790 (''.''))) then (case ((string_drop s35790 ((string_length (''.''))))) of s35800 => (case ((reg_name_matches_prefix s35800 :: (( 5 Word.word * ii)) option)) of Some ((rd, s35810)) => (let s35820 = (string_drop s35800 s35810) in if ((string_startswith s35820 (''.''))) then (case ((string_drop s35820 ((string_length (''.''))))) of s35830 => (case ((hex_bits_4_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s35830 :: (( 4 Word.word * ii)) option)) of Some ((fm, s35840)) => (let p00 = (string_drop s35830 s35840) in if (((p00 = ('''')))) then Some (pred, succ, rs, rd, fm) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None))\ for s35690 :: " string " \ \\val _s3562_ : string -> maybe (mword ty3)\\ definition s3562 :: \ string \((3)Word.word)option \ where \ s3562 s35630 = ( (let s35640 = s35630 in if ((string_startswith s35640 (''c.srai.hint.''))) then (case ((string_drop s35640 ((string_length (''c.srai.hint.''))))) of s35650 => (case ((creg_name_matches_prefix s35650 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s35660)) => (let p00 = (string_drop s35650 s35660) in if (((p00 = ('''')))) then Some rsd else None) | _ => None ) ) else None))\ for s35630 :: " string " \ \\val _s3556_ : string -> maybe (mword ty3)\\ definition s3556 :: \ string \((3)Word.word)option \ where \ s3556 s35570 = ( (let s35580 = s35570 in if ((string_startswith s35580 (''c.srli.hint.''))) then (case ((string_drop s35580 ((string_length (''c.srli.hint.''))))) of s35590 => (case ((creg_name_matches_prefix s35590 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s35600)) => (let p00 = (string_drop s35590 s35600) in if (((p00 = ('''')))) then Some rsd else None) | _ => None ) ) else None))\ for s35570 :: " string " \ \\val _s3547_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3547 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3547 s35480 = ( (let s35490 = s35480 in if ((string_startswith s35490 (''c.slli.hint.''))) then (case ((string_drop s35490 ((string_length (''c.slli.hint.''))))) of s35500 => (case ((reg_name_matches_prefix s35500 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s35510)) => (let s35520 = (string_drop s35500 s35510) in if ((string_startswith s35520 (''.''))) then (case ((string_drop s35520 ((string_length (''.''))))) of s35530 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s35530 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s35540)) => (let p00 = (string_drop s35530 s35540) in if (((p00 = ('''')))) then Some (rsd, shamt) else None) | _ => None ) ) else None) | _ => None ) ) else None))\ for s35480 :: " string " \ \\val _s3541_ : string -> maybe (mword ty5)\\ definition s3541 :: \ string \((5)Word.word)option \ where \ s3541 s35420 = ( (let s35430 = s35420 in if ((string_startswith s35430 (''c.add.hint.''))) then (case ((string_drop s35430 ((string_length (''c.add.hint.''))))) of s35440 => (case ((reg_name_matches_prefix s35440 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s35450)) => (let p00 = (string_drop s35440 s35450) in if (((p00 = ('''')))) then Some rs2 else None) | _ => None ) ) else None))\ for s35420 :: " string " \ \\val _s3535_ : string -> maybe (mword ty5)\\ definition s3535 :: \ string \((5)Word.word)option \ where \ s3535 s35360 = ( (let s35370 = s35360 in if ((string_startswith s35370 (''c.mv.hint.''))) then (case ((string_drop s35370 ((string_length (''c.mv.hint.''))))) of s35380 => (case ((reg_name_matches_prefix s35380 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s35390)) => (let p00 = (string_drop s35380 s35390) in if (((p00 = ('''')))) then Some rs2 else None) | _ => None ) ) else None))\ for s35360 :: " string " \ \\val _s3529_ : string -> maybe (mword ty6)\\ definition s3529 :: \ string \((6)Word.word)option \ where \ s3529 s35300 = ( (let s35310 = s35300 in if ((string_startswith s35310 (''c.lui.hint.''))) then (case ((string_drop s35310 ((string_length (''c.lui.hint.''))))) of s35320 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s35320 :: (( 6 Word.word * ii)) option)) of Some ((imm, s35330)) => (let p00 = (string_drop s35320 s35330) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) else None))\ for s35300 :: " string " \ \\val _s3523_ : string -> maybe (mword ty6)\\ definition s3523 :: \ string \((6)Word.word)option \ where \ s3523 s35240 = ( (let s35250 = s35240 in if ((string_startswith s35250 (''c.li.hint.''))) then (case ((string_drop s35250 ((string_length (''c.li.hint.''))))) of s35260 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s35260 :: (( 6 Word.word * ii)) option)) of Some ((imm, s35270)) => (let p00 = (string_drop s35260 s35270) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) else None))\ for s35240 :: " string " \ \\val _s3517_ : string -> maybe (mword ty5)\\ definition s3517 :: \ string \((5)Word.word)option \ where \ s3517 s35180 = ( (let s35190 = s35180 in if ((string_startswith s35190 (''c.addi.hint.''))) then (case ((string_drop s35190 ((string_length (''c.addi.hint.''))))) of s35200 => (case ((reg_name_matches_prefix s35200 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s35210)) => (let p00 = (string_drop s35200 s35210) in if (((p00 = ('''')))) then Some rsd else None) | _ => None ) ) else None))\ for s35180 :: " string " \ \\val _s3511_ : string -> maybe (mword ty6)\\ definition s3511 :: \ string \((6)Word.word)option \ where \ s3511 s35120 = ( (let s35130 = s35120 in if ((string_startswith s35130 (''c.nop.hint.''))) then (case ((string_drop s35130 ((string_length (''c.nop.hint.''))))) of s35140 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s35140 :: (( 6 Word.word * ii)) option)) of Some ((imm, s35150)) => (let p00 = (string_drop s35140 s35150) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) else None))\ for s35120 :: " string " \ \\val _s3494_ : string -> maybe ((csrop * mword ty5 * mword ty12 * mword ty5))\\ definition s3494 :: \ string \(csrop*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s3494 s34960 = ( (case ((csr_mnemonic_matches_prefix s34960)) of Some ((op1, s34970)) => (case ((string_drop s34960 s34970)) of s34980 => (case ((spc_matches_prefix0 s34980)) of Some ((_, s34990)) => (case ((string_drop s34980 s34990)) of s35000 => (case ((reg_name_matches_prefix s35000 :: (( 5 Word.word * ii)) option)) of Some ((rd, s35010)) => (case ((string_drop s35000 s35010)) of s35020 => (case ((sep_matches_prefix s35020)) of Some ((_, s35030)) => (case ((string_drop s35020 s35030)) of s35040 => (case ((csr_name_map_matches_prefix s35040 :: (( 12 Word.word * ii)) option)) of Some ((csr, s35050)) => (case ((string_drop s35040 s35050)) of s35060 => (case ((sep_matches_prefix s35060)) of Some ((_, s35070)) => (case ((string_drop s35060 s35070)) of s35080 => (case ((reg_name_matches_prefix s35080 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s35090)) => (let p00 = (string_drop s35080 s35090) in if (((p00 = ('''')))) then Some (op1, rd, csr, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s34960 :: " string " \ \\val _s3476_ : string -> maybe ((csrop * mword ty5 * mword ty12 * mword ty5))\\ definition s3476 :: \ string \(csrop*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s3476 s34780 = ( (case ((csr_mnemonic_matches_prefix s34780)) of Some ((op1, s34790)) => (let s34800 = (string_drop s34780 s34790) in if ((string_startswith s34800 (''i''))) then (case ((string_drop s34800 ((string_length (''i''))))) of s34810 => (case ((spc_matches_prefix0 s34810)) of Some ((_, s34820)) => (case ((string_drop s34810 s34820)) of s34830 => (case ((reg_name_matches_prefix s34830 :: (( 5 Word.word * ii)) option)) of Some ((rd, s34840)) => (case ((string_drop s34830 s34840)) of s34850 => (case ((sep_matches_prefix s34850)) of Some ((_, s34860)) => (case ((string_drop s34850 s34860)) of s34870 => (case ((csr_name_map_matches_prefix s34870 :: (( 12 Word.word * ii)) option)) of Some ((csr, s34880)) => (case ((string_drop s34870 s34880)) of s34890 => (case ((sep_matches_prefix s34890)) of Some ((_, s34900)) => (case ((string_drop s34890 s34900)) of s34910 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s34910 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s34920)) => (let p00 = (string_drop s34910 s34920) in if (((p00 = ('''')))) then Some (op1, rd, csr, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ))\ for s34780 :: " string " \ \\val _s3457_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5))\\ definition s3457 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3457 s34580 = ( (let s34590 = s34580 in if ((string_startswith s34590 (''rem''))) then (case ((string_drop s34590 ((string_length (''rem''))))) of s34600 => (case ((maybe_not_u_matches_prefix s34600)) of Some ((s, s34610)) => (let s34620 = (string_drop s34600 s34610) in if ((string_startswith s34620 (''w''))) then (case ((string_drop s34620 ((string_length (''w''))))) of s34630 => (case ((spc_matches_prefix0 s34630)) of Some ((_, s34640)) => (case ((string_drop s34630 s34640)) of s34650 => (case ((reg_name_matches_prefix s34650 :: (( 5 Word.word * ii)) option)) of Some ((rd, s34660)) => (case ((string_drop s34650 s34660)) of s34670 => (case ((sep_matches_prefix s34670)) of Some ((_, s34680)) => (case ((string_drop s34670 s34680)) of s34690 => (case ((reg_name_matches_prefix s34690 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s34700)) => (case ((string_drop s34690 s34700)) of s34710 => (case ((sep_matches_prefix s34710)) of Some ((_, s34720)) => (case ((string_drop s34710 s34720)) of s34730 => (case ((reg_name_matches_prefix s34730 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s34740)) => (let p00 = (string_drop s34730 s34740) in if (((p00 = ('''')))) then Some (s, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) else None))\ for s34580 :: " string " \ \\val _s3438_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5))\\ definition s3438 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3438 s34390 = ( (let s34400 = s34390 in if ((string_startswith s34400 (''div''))) then (case ((string_drop s34400 ((string_length (''div''))))) of s34410 => (case ((maybe_not_u_matches_prefix s34410)) of Some ((s, s34420)) => (let s34430 = (string_drop s34410 s34420) in if ((string_startswith s34430 (''w''))) then (case ((string_drop s34430 ((string_length (''w''))))) of s34440 => (case ((spc_matches_prefix0 s34440)) of Some ((_, s34450)) => (case ((string_drop s34440 s34450)) of s34460 => (case ((reg_name_matches_prefix s34460 :: (( 5 Word.word * ii)) option)) of Some ((rd, s34470)) => (case ((string_drop s34460 s34470)) of s34480 => (case ((sep_matches_prefix s34480)) of Some ((_, s34490)) => (case ((string_drop s34480 s34490)) of s34500 => (case ((reg_name_matches_prefix s34500 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s34510)) => (case ((string_drop s34500 s34510)) of s34520 => (case ((sep_matches_prefix s34520)) of Some ((_, s34530)) => (case ((string_drop s34520 s34530)) of s34540 => (case ((reg_name_matches_prefix s34540 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s34550)) => (let p00 = (string_drop s34540 s34550) in if (((p00 = ('''')))) then Some (s, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) else None))\ for s34390 :: " string " \ \\val _s3422_ : string -> maybe ((mword ty5 * mword ty5 * mword ty5))\\ definition s3422 :: \ string \((5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3422 s34230 = ( (let s34240 = s34230 in if ((string_startswith s34240 (''mulw''))) then (case ((string_drop s34240 ((string_length (''mulw''))))) of s34250 => (case ((spc_matches_prefix0 s34250)) of Some ((_, s34260)) => (case ((string_drop s34250 s34260)) of s34270 => (case ((reg_name_matches_prefix s34270 :: (( 5 Word.word * ii)) option)) of Some ((rd, s34280)) => (case ((string_drop s34270 s34280)) of s34290 => (case ((sep_matches_prefix s34290)) of Some ((_, s34300)) => (case ((string_drop s34290 s34300)) of s34310 => (case ((reg_name_matches_prefix s34310 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s34320)) => (case ((string_drop s34310 s34320)) of s34330 => (case ((sep_matches_prefix s34330)) of Some ((_, s34340)) => (case ((string_drop s34330 s34340)) of s34350 => (case ((reg_name_matches_prefix s34350 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s34360)) => (let p00 = (string_drop s34350 s34360) in if (((p00 = ('''')))) then Some (rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s34230 :: " string " \ \\val _s3404_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5))\\ definition s3404 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3404 s34050 = ( (let s34060 = s34050 in if ((string_startswith s34060 (''rem''))) then (case ((string_drop s34060 ((string_length (''rem''))))) of s34070 => (case ((maybe_not_u_matches_prefix s34070)) of Some ((s, s34080)) => (case ((string_drop s34070 s34080)) of s34090 => (case ((spc_matches_prefix0 s34090)) of Some ((_, s34100)) => (case ((string_drop s34090 s34100)) of s34110 => (case ((reg_name_matches_prefix s34110 :: (( 5 Word.word * ii)) option)) of Some ((rd, s34120)) => (case ((string_drop s34110 s34120)) of s34130 => (case ((sep_matches_prefix s34130)) of Some ((_, s34140)) => (case ((string_drop s34130 s34140)) of s34150 => (case ((reg_name_matches_prefix s34150 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s34160)) => (case ((string_drop s34150 s34160)) of s34170 => (case ((sep_matches_prefix s34170)) of Some ((_, s34180)) => (case ((string_drop s34170 s34180)) of s34190 => (case ((reg_name_matches_prefix s34190 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s34200)) => (let p00 = (string_drop s34190 s34200) in if (((p00 = ('''')))) then Some (s, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s34050 :: " string " \ \\val _s3386_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5))\\ definition s3386 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3386 s33870 = ( (let s33880 = s33870 in if ((string_startswith s33880 (''div''))) then (case ((string_drop s33880 ((string_length (''div''))))) of s33890 => (case ((maybe_not_u_matches_prefix s33890)) of Some ((s, s33900)) => (case ((string_drop s33890 s33900)) of s33910 => (case ((spc_matches_prefix0 s33910)) of Some ((_, s33920)) => (case ((string_drop s33910 s33920)) of s33930 => (case ((reg_name_matches_prefix s33930 :: (( 5 Word.word * ii)) option)) of Some ((rd, s33940)) => (case ((string_drop s33930 s33940)) of s33950 => (case ((sep_matches_prefix s33950)) of Some ((_, s33960)) => (case ((string_drop s33950 s33960)) of s33970 => (case ((reg_name_matches_prefix s33970 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s33980)) => (case ((string_drop s33970 s33980)) of s33990 => (case ((sep_matches_prefix s33990)) of Some ((_, s34000)) => (case ((string_drop s33990 s34000)) of s34010 => (case ((reg_name_matches_prefix s34010 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s34020)) => (let p00 = (string_drop s34010 s34020) in if (((p00 = ('''')))) then Some (s, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s33870 :: " string " \ \\val _s3369_ : string -> maybe ((bool * bool * bool * mword ty5 * mword ty5 * mword ty5))\\ definition s3369 :: \ string \(bool*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s3369 s33710 = ( (case ((mul_mnemonic_matches_prefix s33710)) of Some (((high, signed1, signed2), s33720)) => (case ((string_drop s33710 s33720)) of s33730 => (case ((spc_matches_prefix0 s33730)) of Some ((_, s33740)) => (case ((string_drop s33730 s33740)) of s33750 => (case ((reg_name_matches_prefix s33750 :: (( 5 Word.word * ii)) option)) of Some ((rd, s33760)) => (case ((string_drop s33750 s33760)) of s33770 => (case ((sep_matches_prefix s33770)) of Some ((_, s33780)) => (case ((string_drop s33770 s33780)) of s33790 => (case ((reg_name_matches_prefix s33790 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s33800)) => (case ((string_drop s33790 s33800)) of s33810 => (case ((sep_matches_prefix s33810)) of Some ((_, s33820)) => (case ((string_drop s33810 s33820)) of s33830 => (case ((reg_name_matches_prefix s33830 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s33840)) => (let p00 = (string_drop s33830 s33840) in if (((p00 = ('''')))) then Some (high, signed1, signed2, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s33710 :: " string " \ \\val _s3357_ : string -> maybe ((mword ty5 * mword ty5))\\ definition s3357 :: \ string \((5)Word.word*(5)Word.word)option \ where \ s3357 s33580 = ( (let s33590 = s33580 in if ((string_startswith s33590 (''c.add''))) then (case ((string_drop s33590 ((string_length (''c.add''))))) of s33600 => (case ((spc_matches_prefix0 s33600)) of Some ((_, s33610)) => (case ((string_drop s33600 s33610)) of s33620 => (case ((reg_name_matches_prefix s33620 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s33630)) => (case ((string_drop s33620 s33630)) of s33640 => (case ((sep_matches_prefix s33640)) of Some ((_, s33650)) => (case ((string_drop s33640 s33650)) of s33660 => (case ((reg_name_matches_prefix s33660 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s33670)) => (let p00 = (string_drop s33660 s33670) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s33580 :: " string " \ \\val _s3345_ : string -> maybe ((mword ty5 * mword ty5))\\ definition s3345 :: \ string \((5)Word.word*(5)Word.word)option \ where \ s3345 s33460 = ( (let s33470 = s33460 in if ((string_startswith s33470 (''c.mv''))) then (case ((string_drop s33470 ((string_length (''c.mv''))))) of s33480 => (case ((spc_matches_prefix0 s33480)) of Some ((_, s33490)) => (case ((string_drop s33480 s33490)) of s33500 => (case ((reg_name_matches_prefix s33500 :: (( 5 Word.word * ii)) option)) of Some ((rd, s33510)) => (case ((string_drop s33500 s33510)) of s33520 => (case ((sep_matches_prefix s33520)) of Some ((_, s33530)) => (case ((string_drop s33520 s33530)) of s33540 => (case ((reg_name_matches_prefix s33540 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s33550)) => (let p00 = (string_drop s33540 s33550) in if (((p00 = ('''')))) then Some (rd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s33460 :: " string " \ \\val _s3337_ : string -> maybe (mword ty5)\\ definition s3337 :: \ string \((5)Word.word)option \ where \ s3337 s33380 = ( (let s33390 = s33380 in if ((string_startswith s33390 (''c.jalr''))) then (case ((string_drop s33390 ((string_length (''c.jalr''))))) of s33400 => (case ((spc_matches_prefix0 s33400)) of Some ((_, s33410)) => (case ((string_drop s33400 s33410)) of s33420 => (case ((reg_name_matches_prefix s33420 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s33430)) => (let p00 = (string_drop s33420 s33430) in if (((p00 = ('''')))) then Some rs1 else None) | _ => None ) ) | _ => None ) ) else None))\ for s33380 :: " string " \ \\val _s3329_ : string -> maybe (mword ty5)\\ definition s3329 :: \ string \((5)Word.word)option \ where \ s3329 s33300 = ( (let s33310 = s33300 in if ((string_startswith s33310 (''c.jr''))) then (case ((string_drop s33310 ((string_length (''c.jr''))))) of s33320 => (case ((spc_matches_prefix0 s33320)) of Some ((_, s33330)) => (case ((string_drop s33320 s33330)) of s33340 => (case ((reg_name_matches_prefix s33340 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s33350)) => (let p00 = (string_drop s33340 s33350) in if (((p00 = ('''')))) then Some rs1 else None) | _ => None ) ) | _ => None ) ) else None))\ for s33300 :: " string " \ \\val _s3317_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3317 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3317 s33180 = ( (let s33190 = s33180 in if ((string_startswith s33190 (''c.sdsp''))) then (case ((string_drop s33190 ((string_length (''c.sdsp''))))) of s33200 => (case ((spc_matches_prefix0 s33200)) of Some ((_, s33210)) => (case ((string_drop s33200 s33210)) of s33220 => (case ((reg_name_matches_prefix s33220 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s33230)) => (case ((string_drop s33220 s33230)) of s33240 => (case ((sep_matches_prefix s33240)) of Some ((_, s33250)) => (case ((string_drop s33240 s33250)) of s33260 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s33260 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s33270)) => (let p00 = (string_drop s33260 s33270) in if (((p00 = ('''')))) then Some (rs2, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s33180 :: " string " \ \\val _s3305_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3305 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3305 s33060 = ( (let s33070 = s33060 in if ((string_startswith s33070 (''c.swsp''))) then (case ((string_drop s33070 ((string_length (''c.swsp''))))) of s33080 => (case ((spc_matches_prefix0 s33080)) of Some ((_, s33090)) => (case ((string_drop s33080 s33090)) of s33100 => (case ((reg_name_matches_prefix s33100 :: (( 5 Word.word * ii)) option)) of Some ((rd, s33110)) => (case ((string_drop s33100 s33110)) of s33120 => (case ((sep_matches_prefix s33120)) of Some ((_, s33130)) => (case ((string_drop s33120 s33130)) of s33140 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s33140 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s33150)) => (let p00 = (string_drop s33140 s33150) in if (((p00 = ('''')))) then Some (rd, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s33060 :: " string " \ \\val _s3293_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3293 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3293 s32940 = ( (let s32950 = s32940 in if ((string_startswith s32950 (''c.ldsp''))) then (case ((string_drop s32950 ((string_length (''c.ldsp''))))) of s32960 => (case ((spc_matches_prefix0 s32960)) of Some ((_, s32970)) => (case ((string_drop s32960 s32970)) of s32980 => (case ((reg_name_matches_prefix s32980 :: (( 5 Word.word * ii)) option)) of Some ((rd, s32990)) => (case ((string_drop s32980 s32990)) of s33000 => (case ((sep_matches_prefix s33000)) of Some ((_, s33010)) => (case ((string_drop s33000 s33010)) of s33020 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s33020 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s33030)) => (let p00 = (string_drop s33020 s33030) in if (((p00 = ('''')))) then Some (rd, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s32940 :: " string " \ \\val _s3281_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3281 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3281 s32820 = ( (let s32830 = s32820 in if ((string_startswith s32830 (''c.lwsp''))) then (case ((string_drop s32830 ((string_length (''c.lwsp''))))) of s32840 => (case ((spc_matches_prefix0 s32840)) of Some ((_, s32850)) => (case ((string_drop s32840 s32850)) of s32860 => (case ((reg_name_matches_prefix s32860 :: (( 5 Word.word * ii)) option)) of Some ((rd, s32870)) => (case ((string_drop s32860 s32870)) of s32880 => (case ((sep_matches_prefix s32880)) of Some ((_, s32890)) => (case ((string_drop s32880 s32890)) of s32900 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s32900 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s32910)) => (let p00 = (string_drop s32900 s32910) in if (((p00 = ('''')))) then Some (rd, uimm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s32820 :: " string " \ \\val _s3269_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3269 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3269 s32700 = ( (let s32710 = s32700 in if ((string_startswith s32710 (''c.slli''))) then (case ((string_drop s32710 ((string_length (''c.slli''))))) of s32720 => (case ((spc_matches_prefix0 s32720)) of Some ((_, s32730)) => (case ((string_drop s32720 s32730)) of s32740 => (case ((reg_name_matches_prefix s32740 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s32750)) => (case ((string_drop s32740 s32750)) of s32760 => (case ((sep_matches_prefix s32760)) of Some ((_, s32770)) => (case ((string_drop s32760 s32770)) of s32780 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s32780 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s32790)) => (let p00 = (string_drop s32780 s32790) in if (((p00 = ('''')))) then Some (rsd, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s32700 :: " string " \ \\val _s3257_ : string -> maybe ((mword ty3 * mword ty8))\\ definition s3257 :: \ string \((3)Word.word*(8)Word.word)option \ where \ s3257 s32580 = ( (let s32590 = s32580 in if ((string_startswith s32590 (''c.bnez''))) then (case ((string_drop s32590 ((string_length (''c.bnez''))))) of s32600 => (case ((spc_matches_prefix0 s32600)) of Some ((_, s32610)) => (case ((string_drop s32600 s32610)) of s32620 => (case ((creg_name_matches_prefix s32620 :: (( 3 Word.word * ii)) option)) of Some ((rs, s32630)) => (case ((string_drop s32620 s32630)) of s32640 => (case ((sep_matches_prefix s32640)) of Some ((_, s32650)) => (case ((string_drop s32640 s32650)) of s32660 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s32660 :: (( 8 Word.word * ii)) option)) of Some ((imm, s32670)) => (let p00 = (string_drop s32660 s32670) in if (((p00 = ('''')))) then Some (rs, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s32580 :: " string " \ \\val _s3245_ : string -> maybe ((mword ty3 * mword ty8))\\ definition s3245 :: \ string \((3)Word.word*(8)Word.word)option \ where \ s3245 s32460 = ( (let s32470 = s32460 in if ((string_startswith s32470 (''c.beqz''))) then (case ((string_drop s32470 ((string_length (''c.beqz''))))) of s32480 => (case ((spc_matches_prefix0 s32480)) of Some ((_, s32490)) => (case ((string_drop s32480 s32490)) of s32500 => (case ((creg_name_matches_prefix s32500 :: (( 3 Word.word * ii)) option)) of Some ((rs, s32510)) => (case ((string_drop s32500 s32510)) of s32520 => (case ((sep_matches_prefix s32520)) of Some ((_, s32530)) => (case ((string_drop s32520 s32530)) of s32540 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s32540 :: (( 8 Word.word * ii)) option)) of Some ((imm, s32550)) => (let p00 = (string_drop s32540 s32550) in if (((p00 = ('''')))) then Some (rs, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s32460 :: " string " \ \\val _s3237_ : string -> maybe (mword ty11)\\ definition s3237 :: \ string \((11)Word.word)option \ where \ s3237 s32380 = ( (let s32390 = s32380 in if ((string_startswith s32390 (''c.j''))) then (case ((string_drop s32390 ((string_length (''c.j''))))) of s32400 => (case ((spc_matches_prefix0 s32400)) of Some ((_, s32410)) => (case ((string_drop s32400 s32410)) of s32420 => (case ((hex_bits_11_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s32420 :: (( 11 Word.word * ii)) option)) of Some ((imm, s32430)) => (let p00 = (string_drop s32420 s32430) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) | _ => None ) ) else None))\ for s32380 :: " string " \ \\val _s3225_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s3225 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s3225 s32260 = ( (let s32270 = s32260 in if ((string_startswith s32270 (''c.addw''))) then (case ((string_drop s32270 ((string_length (''c.addw''))))) of s32280 => (case ((spc_matches_prefix0 s32280)) of Some ((_, s32290)) => (case ((string_drop s32280 s32290)) of s32300 => (case ((creg_name_matches_prefix s32300 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s32310)) => (case ((string_drop s32300 s32310)) of s32320 => (case ((sep_matches_prefix s32320)) of Some ((_, s32330)) => (case ((string_drop s32320 s32330)) of s32340 => (case ((creg_name_matches_prefix s32340 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s32350)) => (let p00 = (string_drop s32340 s32350) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s32260 :: " string " \ \\val _s3213_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s3213 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s3213 s32140 = ( (let s32150 = s32140 in if ((string_startswith s32150 (''c.subw''))) then (case ((string_drop s32150 ((string_length (''c.subw''))))) of s32160 => (case ((spc_matches_prefix0 s32160)) of Some ((_, s32170)) => (case ((string_drop s32160 s32170)) of s32180 => (case ((creg_name_matches_prefix s32180 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s32190)) => (case ((string_drop s32180 s32190)) of s32200 => (case ((sep_matches_prefix s32200)) of Some ((_, s32210)) => (case ((string_drop s32200 s32210)) of s32220 => (case ((creg_name_matches_prefix s32220 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s32230)) => (let p00 = (string_drop s32220 s32230) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s32140 :: " string " \ \\val _s3201_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s3201 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s3201 s32020 = ( (let s32030 = s32020 in if ((string_startswith s32030 (''c.and''))) then (case ((string_drop s32030 ((string_length (''c.and''))))) of s32040 => (case ((spc_matches_prefix0 s32040)) of Some ((_, s32050)) => (case ((string_drop s32040 s32050)) of s32060 => (case ((creg_name_matches_prefix s32060 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s32070)) => (case ((string_drop s32060 s32070)) of s32080 => (case ((sep_matches_prefix s32080)) of Some ((_, s32090)) => (case ((string_drop s32080 s32090)) of s32100 => (case ((creg_name_matches_prefix s32100 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s32110)) => (let p00 = (string_drop s32100 s32110) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s32020 :: " string " \ \\val _s3189_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s3189 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s3189 s31900 = ( (let s31910 = s31900 in if ((string_startswith s31910 (''c.or''))) then (case ((string_drop s31910 ((string_length (''c.or''))))) of s31920 => (case ((spc_matches_prefix0 s31920)) of Some ((_, s31930)) => (case ((string_drop s31920 s31930)) of s31940 => (case ((creg_name_matches_prefix s31940 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s31950)) => (case ((string_drop s31940 s31950)) of s31960 => (case ((sep_matches_prefix s31960)) of Some ((_, s31970)) => (case ((string_drop s31960 s31970)) of s31980 => (case ((creg_name_matches_prefix s31980 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s31990)) => (let p00 = (string_drop s31980 s31990) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s31900 :: " string " \ \\val _s3177_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s3177 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s3177 s31780 = ( (let s31790 = s31780 in if ((string_startswith s31790 (''c.xor''))) then (case ((string_drop s31790 ((string_length (''c.xor''))))) of s31800 => (case ((spc_matches_prefix0 s31800)) of Some ((_, s31810)) => (case ((string_drop s31800 s31810)) of s31820 => (case ((creg_name_matches_prefix s31820 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s31830)) => (case ((string_drop s31820 s31830)) of s31840 => (case ((sep_matches_prefix s31840)) of Some ((_, s31850)) => (case ((string_drop s31840 s31850)) of s31860 => (case ((creg_name_matches_prefix s31860 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s31870)) => (let p00 = (string_drop s31860 s31870) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s31780 :: " string " \ \\val _s3165_ : string -> maybe ((mword ty3 * mword ty3))\\ definition s3165 :: \ string \((3)Word.word*(3)Word.word)option \ where \ s3165 s31660 = ( (let s31670 = s31660 in if ((string_startswith s31670 (''c.sub''))) then (case ((string_drop s31670 ((string_length (''c.sub''))))) of s31680 => (case ((spc_matches_prefix0 s31680)) of Some ((_, s31690)) => (case ((string_drop s31680 s31690)) of s31700 => (case ((creg_name_matches_prefix s31700 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s31710)) => (case ((string_drop s31700 s31710)) of s31720 => (case ((sep_matches_prefix s31720)) of Some ((_, s31730)) => (case ((string_drop s31720 s31730)) of s31740 => (case ((creg_name_matches_prefix s31740 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s31750)) => (let p00 = (string_drop s31740 s31750) in if (((p00 = ('''')))) then Some (rsd, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s31660 :: " string " \ \\val _s3153_ : string -> maybe ((mword ty3 * mword ty6))\\ definition s3153 :: \ string \((3)Word.word*(6)Word.word)option \ where \ s3153 s31540 = ( (let s31550 = s31540 in if ((string_startswith s31550 (''c.andi''))) then (case ((string_drop s31550 ((string_length (''c.andi''))))) of s31560 => (case ((spc_matches_prefix0 s31560)) of Some ((_, s31570)) => (case ((string_drop s31560 s31570)) of s31580 => (case ((creg_name_matches_prefix s31580 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s31590)) => (case ((string_drop s31580 s31590)) of s31600 => (case ((sep_matches_prefix s31600)) of Some ((_, s31610)) => (case ((string_drop s31600 s31610)) of s31620 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s31620 :: (( 6 Word.word * ii)) option)) of Some ((imm, s31630)) => (let p00 = (string_drop s31620 s31630) in if (((p00 = ('''')))) then Some (rsd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s31540 :: " string " \ \\val _s3141_ : string -> maybe ((mword ty3 * mword ty6))\\ definition s3141 :: \ string \((3)Word.word*(6)Word.word)option \ where \ s3141 s31420 = ( (let s31430 = s31420 in if ((string_startswith s31430 (''c.srai''))) then (case ((string_drop s31430 ((string_length (''c.srai''))))) of s31440 => (case ((spc_matches_prefix0 s31440)) of Some ((_, s31450)) => (case ((string_drop s31440 s31450)) of s31460 => (case ((creg_name_matches_prefix s31460 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s31470)) => (case ((string_drop s31460 s31470)) of s31480 => (case ((sep_matches_prefix s31480)) of Some ((_, s31490)) => (case ((string_drop s31480 s31490)) of s31500 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s31500 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s31510)) => (let p00 = (string_drop s31500 s31510) in if (((p00 = ('''')))) then Some (rsd, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s31420 :: " string " \ \\val _s3129_ : string -> maybe ((mword ty3 * mword ty6))\\ definition s3129 :: \ string \((3)Word.word*(6)Word.word)option \ where \ s3129 s31300 = ( (let s31310 = s31300 in if ((string_startswith s31310 (''c.srli''))) then (case ((string_drop s31310 ((string_length (''c.srli''))))) of s31320 => (case ((spc_matches_prefix0 s31320)) of Some ((_, s31330)) => (case ((string_drop s31320 s31330)) of s31340 => (case ((creg_name_matches_prefix s31340 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s31350)) => (case ((string_drop s31340 s31350)) of s31360 => (case ((sep_matches_prefix s31360)) of Some ((_, s31370)) => (case ((string_drop s31360 s31370)) of s31380 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s31380 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s31390)) => (let p00 = (string_drop s31380 s31390) in if (((p00 = ('''')))) then Some (rsd, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s31300 :: " string " \ \\val _s3117_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3117 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3117 s31180 = ( (let s31190 = s31180 in if ((string_startswith s31190 (''c.lui''))) then (case ((string_drop s31190 ((string_length (''c.lui''))))) of s31200 => (case ((spc_matches_prefix0 s31200)) of Some ((_, s31210)) => (case ((string_drop s31200 s31210)) of s31220 => (case ((reg_name_matches_prefix s31220 :: (( 5 Word.word * ii)) option)) of Some ((rd, s31230)) => (case ((string_drop s31220 s31230)) of s31240 => (case ((sep_matches_prefix s31240)) of Some ((_, s31250)) => (case ((string_drop s31240 s31250)) of s31260 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s31260 :: (( 6 Word.word * ii)) option)) of Some ((imm, s31270)) => (let p00 = (string_drop s31260 s31270) in if (((p00 = ('''')))) then Some (rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s31180 :: " string " \ \\val _s3109_ : string -> maybe (mword ty6)\\ definition s3109 :: \ string \((6)Word.word)option \ where \ s3109 s31100 = ( (let s31110 = s31100 in if ((string_startswith s31110 (''c.addi16sp''))) then (case ((string_drop s31110 ((string_length (''c.addi16sp''))))) of s31120 => (case ((spc_matches_prefix0 s31120)) of Some ((_, s31130)) => (case ((string_drop s31120 s31130)) of s31140 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s31140 :: (( 6 Word.word * ii)) option)) of Some ((imm, s31150)) => (let p00 = (string_drop s31140 s31150) in if (((p00 = ('''')))) then Some imm else None) | _ => None ) ) | _ => None ) ) else None))\ for s31100 :: " string " \ \\val _s3097_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3097 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3097 s30980 = ( (let s30990 = s30980 in if ((string_startswith s30990 (''c.li''))) then (case ((string_drop s30990 ((string_length (''c.li''))))) of s31000 => (case ((spc_matches_prefix0 s31000)) of Some ((_, s31010)) => (case ((string_drop s31000 s31010)) of s31020 => (case ((reg_name_matches_prefix s31020 :: (( 5 Word.word * ii)) option)) of Some ((rd, s31030)) => (case ((string_drop s31020 s31030)) of s31040 => (case ((sep_matches_prefix s31040)) of Some ((_, s31050)) => (case ((string_drop s31040 s31050)) of s31060 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s31060 :: (( 6 Word.word * ii)) option)) of Some ((imm, s31070)) => (let p00 = (string_drop s31060 s31070) in if (((p00 = ('''')))) then Some (rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s30980 :: " string " \ \\val _s3085_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3085 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3085 s30860 = ( (let s30870 = s30860 in if ((string_startswith s30870 (''c.addiw''))) then (case ((string_drop s30870 ((string_length (''c.addiw''))))) of s30880 => (case ((spc_matches_prefix0 s30880)) of Some ((_, s30890)) => (case ((string_drop s30880 s30890)) of s30900 => (case ((reg_name_matches_prefix s30900 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s30910)) => (case ((string_drop s30900 s30910)) of s30920 => (case ((sep_matches_prefix s30920)) of Some ((_, s30930)) => (case ((string_drop s30920 s30930)) of s30940 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s30940 :: (( 6 Word.word * ii)) option)) of Some ((imm, s30950)) => (let p00 = (string_drop s30940 s30950) in if (((p00 = ('''')))) then Some (rsd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s30860 :: " string " \ \\val _s3077_ : string -> maybe (mword ty11)\\ definition s3077 :: \ string \((11)Word.word)option \ where \ s3077 s30780 = ( (let s30790 = s30780 in if ((string_startswith s30790 (''c.jal''))) then (case ((string_drop s30790 ((string_length (''c.jal''))))) of s30800 => (case ((spc_matches_prefix0 s30800)) of Some ((_, s30810)) => (case ((string_drop s30800 s30810)) of s30820 => (case ((hex_bits_12_matches_prefix0 s30820 :: (( 12 Word.word * ii)) option)) of Some ((v__1200, s30830)) => if (((((subrange_vec_dec v__1200 (( 0 :: int):: ii) (( 0 :: int):: ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then (let (imm :: 11 Word.word) = ((subrange_vec_dec v__1200 (( 11 :: int):: ii) (( 1 :: int):: ii) :: 11 Word.word)) in (let (imm :: 11 Word.word) = ((subrange_vec_dec v__1200 (( 11 :: int):: ii) (( 1 :: int):: ii) :: 11 Word.word)) in (let p00 = (string_drop s30820 s30830) in if (((p00 = ('''')))) then Some imm else None))) else None | _ => None ) ) | _ => None ) ) else None))\ for s30780 :: " string " \ \\val _s3065_ : string -> maybe ((mword ty5 * mword ty6))\\ definition s3065 :: \ string \((5)Word.word*(6)Word.word)option \ where \ s3065 s30660 = ( (let s30670 = s30660 in if ((string_startswith s30670 (''c.addi''))) then (case ((string_drop s30670 ((string_length (''c.addi''))))) of s30680 => (case ((spc_matches_prefix0 s30680)) of Some ((_, s30690)) => (case ((string_drop s30680 s30690)) of s30700 => (case ((reg_name_matches_prefix s30700 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s30710)) => (case ((string_drop s30700 s30710)) of s30720 => (case ((sep_matches_prefix s30720)) of Some ((_, s30730)) => (case ((string_drop s30720 s30730)) of s30740 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s30740 :: (( 6 Word.word * ii)) option)) of Some ((nzi, s30750)) => (let p00 = (string_drop s30740 s30750) in if (((p00 = ('''')))) then Some (rsd, nzi) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s30660 :: " string " \ \\val _s3049_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s3049 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s3049 s30500 = ( (let s30510 = s30500 in if ((string_startswith s30510 (''c.sd''))) then (case ((string_drop s30510 ((string_length (''c.sd''))))) of s30520 => (case ((spc_matches_prefix0 s30520)) of Some ((_, s30530)) => (case ((string_drop s30520 s30530)) of s30540 => (case ((creg_name_matches_prefix s30540 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s30550)) => (case ((string_drop s30540 s30550)) of s30560 => (case ((sep_matches_prefix s30560)) of Some ((_, s30570)) => (case ((string_drop s30560 s30570)) of s30580 => (case ((creg_name_matches_prefix s30580 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s30590)) => (case ((string_drop s30580 s30590)) of s30600 => (case ((sep_matches_prefix s30600)) of Some ((_, s30610)) => (case ((string_drop s30600 s30610)) of s30620 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s30620 :: (( 8 Word.word * ii)) option)) of Some ((v__1202, s30630)) => if (((((subrange_vec_dec v__1202 (( 2 :: int):: ii) (( 0 :: int):: ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1202 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1202 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s30620 s30630) in if (((p00 = ('''')))) then Some (rsc1, rsc2, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s30500 :: " string " \ \\val _s3033_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s3033 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s3033 s30340 = ( (let s30350 = s30340 in if ((string_startswith s30350 (''c.sw''))) then (case ((string_drop s30350 ((string_length (''c.sw''))))) of s30360 => (case ((spc_matches_prefix0 s30360)) of Some ((_, s30370)) => (case ((string_drop s30360 s30370)) of s30380 => (case ((creg_name_matches_prefix s30380 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s30390)) => (case ((string_drop s30380 s30390)) of s30400 => (case ((sep_matches_prefix s30400)) of Some ((_, s30410)) => (case ((string_drop s30400 s30410)) of s30420 => (case ((creg_name_matches_prefix s30420 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s30430)) => (case ((string_drop s30420 s30430)) of s30440 => (case ((sep_matches_prefix s30440)) of Some ((_, s30450)) => (case ((string_drop s30440 s30450)) of s30460 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s30460 :: (( 7 Word.word * ii)) option)) of Some ((v__1204, s30470)) => if (((((subrange_vec_dec v__1204 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1204 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1204 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s30460 s30470) in if (((p00 = ('''')))) then Some (rsc1, rsc2, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s30340 :: " string " \ \\val _s3017_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s3017 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s3017 s30180 = ( (let s30190 = s30180 in if ((string_startswith s30190 (''c.ld''))) then (case ((string_drop s30190 ((string_length (''c.ld''))))) of s30200 => (case ((spc_matches_prefix0 s30200)) of Some ((_, s30210)) => (case ((string_drop s30200 s30210)) of s30220 => (case ((creg_name_matches_prefix s30220 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s30230)) => (case ((string_drop s30220 s30230)) of s30240 => (case ((sep_matches_prefix s30240)) of Some ((_, s30250)) => (case ((string_drop s30240 s30250)) of s30260 => (case ((creg_name_matches_prefix s30260 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s30270)) => (case ((string_drop s30260 s30270)) of s30280 => (case ((sep_matches_prefix s30280)) of Some ((_, s30290)) => (case ((string_drop s30280 s30290)) of s30300 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s30300 :: (( 8 Word.word * ii)) option)) of Some ((v__1206, s30310)) => if (((((subrange_vec_dec v__1206 (( 2 :: int):: ii) (( 0 :: int):: ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1206 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1206 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s30300 s30310) in if (((p00 = ('''')))) then Some (rdc, rsc, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s30180 :: " string " \ \\val _s3001_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5))\\ definition s3001 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word)option \ where \ s3001 s30020 = ( (let s30030 = s30020 in if ((string_startswith s30030 (''c.lw''))) then (case ((string_drop s30030 ((string_length (''c.lw''))))) of s30040 => (case ((spc_matches_prefix0 s30040)) of Some ((_, s30050)) => (case ((string_drop s30040 s30050)) of s30060 => (case ((creg_name_matches_prefix s30060 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s30070)) => (case ((string_drop s30060 s30070)) of s30080 => (case ((sep_matches_prefix s30080)) of Some ((_, s30090)) => (case ((string_drop s30080 s30090)) of s30100 => (case ((creg_name_matches_prefix s30100 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s30110)) => (case ((string_drop s30100 s30110)) of s30120 => (case ((sep_matches_prefix s30120)) of Some ((_, s30130)) => (case ((string_drop s30120 s30130)) of s30140 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s30140 :: (( 7 Word.word * ii)) option)) of Some ((v__1208, s30150)) => if (((((subrange_vec_dec v__1208 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1208 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1208 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let p00 = (string_drop s30140 s30150) in if (((p00 = ('''')))) then Some (rdc, rsc, uimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s30020 :: " string " \ \\val _s2989_ : string -> maybe ((mword ty3 * mword ty8))\\ definition s2989 :: \ string \((3)Word.word*(8)Word.word)option \ where \ s2989 s29900 = ( (let s29910 = s29900 in if ((string_startswith s29910 (''c.addi4spn''))) then (case ((string_drop s29910 ((string_length (''c.addi4spn''))))) of s29920 => (case ((spc_matches_prefix0 s29920)) of Some ((_, s29930)) => (case ((string_drop s29920 s29930)) of s29940 => (case ((creg_name_matches_prefix s29940 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s29950)) => (case ((string_drop s29940 s29950)) of s29960 => (case ((sep_matches_prefix s29960)) of Some ((_, s29970)) => (case ((string_drop s29960 s29970)) of s29980 => (case ((hex_bits_10_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s29980 :: (( 10 Word.word * ii)) option)) of Some ((v__1210, s29990)) => if (((((subrange_vec_dec v__1210 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (nzimm :: 8 Word.word) = ((subrange_vec_dec v__1210 (( 9 :: int):: ii) (( 2 :: int):: ii) :: 8 Word.word)) in (let (nzimm :: 8 Word.word) = ((subrange_vec_dec v__1210 (( 9 :: int):: ii) (( 2 :: int):: ii) :: 8 Word.word)) in (let p00 = (string_drop s29980 s29990) in if (((p00 = ('''')))) then Some (rdc, nzimm) else None))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s29900 :: " string " \ \\val _s2963_ : string -> maybe ((amoop * word_width * bool * bool * mword ty5 * mword ty5 * mword ty5))\\ definition s2963 :: \ string \(amoop*word_width*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2963 s29650 = ( (case ((amo_mnemonic_matches_prefix s29650)) of Some ((op1, s29660)) => (let s29670 = (string_drop s29650 s29660) in if ((string_startswith s29670 (''.''))) then (case ((string_drop s29670 ((string_length (''.''))))) of s29680 => (case ((size_mnemonic_matches_prefix s29680)) of Some ((width, s29690)) => (case ((string_drop s29680 s29690)) of s29700 => (case ((maybe_aq_matches_prefix s29700)) of Some ((aq, s29710)) => (case ((string_drop s29700 s29710)) of s29720 => (case ((maybe_rl_matches_prefix s29720)) of Some ((rl, s29730)) => (case ((string_drop s29720 s29730)) of s29740 => (case ((spc_matches_prefix0 s29740)) of Some ((_, s29750)) => (case ((string_drop s29740 s29750)) of s29760 => (case ((reg_name_matches_prefix s29760 :: (( 5 Word.word * ii)) option)) of Some ((rd, s29770)) => (case ((string_drop s29760 s29770)) of s29780 => (case ((sep_matches_prefix s29780)) of Some ((_, s29790)) => (case ((string_drop s29780 s29790)) of s29800 => (case ((reg_name_matches_prefix s29800 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s29810)) => (case ((string_drop s29800 s29810)) of s29820 => (case ((sep_matches_prefix s29820)) of Some ((_, s29830)) => (let s29840 = (string_drop s29820 s29830) in if ((string_startswith s29840 (''(''))) then (case ((string_drop s29840 ((string_length (''(''))))) of s29850 => (case ((reg_name_matches_prefix s29850 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s29860)) => (let s29870 = (string_drop s29850 s29860) in if ((string_startswith s29870 ('')''))) then (let p00 = (string_drop s29870 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (op1, width, aq, rl, rd, rs2, rs1) else None) else None) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ))\ for s29650 :: " string " \ \\val _s2941_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty5 * mword ty5))\\ definition s2941 :: \ string \(word_width*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2941 s29420 = ( (let s29430 = s29420 in if ((string_startswith s29430 (''sc.''))) then (case ((string_drop s29430 ((string_length (''sc.''))))) of s29440 => (case ((size_mnemonic_matches_prefix s29440)) of Some ((size1, s29450)) => (case ((string_drop s29440 s29450)) of s29460 => (case ((maybe_aq_matches_prefix s29460)) of Some ((aq, s29470)) => (case ((string_drop s29460 s29470)) of s29480 => (case ((maybe_rl_matches_prefix s29480)) of Some ((rl, s29490)) => (case ((string_drop s29480 s29490)) of s29500 => (case ((spc_matches_prefix0 s29500)) of Some ((_, s29510)) => (case ((string_drop s29500 s29510)) of s29520 => (case ((reg_name_matches_prefix s29520 :: (( 5 Word.word * ii)) option)) of Some ((rd, s29530)) => (case ((string_drop s29520 s29530)) of s29540 => (case ((sep_matches_prefix s29540)) of Some ((_, s29550)) => (case ((string_drop s29540 s29550)) of s29560 => (case ((reg_name_matches_prefix s29560 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s29570)) => (case ((string_drop s29560 s29570)) of s29580 => (case ((sep_matches_prefix s29580)) of Some ((_, s29590)) => (case ((string_drop s29580 s29590)) of s29600 => (case ((reg_name_matches_prefix s29600 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s29610)) => (let p00 = (string_drop s29600 s29610) in if (((p00 = ('''')))) then Some (size1, aq, rl, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s29420 :: " string " \ \\val _s2923_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty5))\\ definition s2923 :: \ string \(word_width*bool*bool*(5)Word.word*(5)Word.word)option \ where \ s2923 s29240 = ( (let s29250 = s29240 in if ((string_startswith s29250 (''lr.''))) then (case ((string_drop s29250 ((string_length (''lr.''))))) of s29260 => (case ((size_mnemonic_matches_prefix s29260)) of Some ((size1, s29270)) => (case ((string_drop s29260 s29270)) of s29280 => (case ((maybe_aq_matches_prefix s29280)) of Some ((aq, s29290)) => (case ((string_drop s29280 s29290)) of s29300 => (case ((maybe_rl_matches_prefix s29300)) of Some ((rl, s29310)) => (case ((string_drop s29300 s29310)) of s29320 => (case ((spc_matches_prefix0 s29320)) of Some ((_, s29330)) => (case ((string_drop s29320 s29330)) of s29340 => (case ((reg_name_matches_prefix s29340 :: (( 5 Word.word * ii)) option)) of Some ((rd, s29350)) => (case ((string_drop s29340 s29350)) of s29360 => (case ((sep_matches_prefix s29360)) of Some ((_, s29370)) => (case ((string_drop s29360 s29370)) of s29380 => (case ((reg_name_matches_prefix s29380 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s29390)) => (let p00 = (string_drop s29380 s29390) in if (((p00 = ('''')))) then Some (size1, aq, rl, rd, rs1) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s29240 :: " string " \ \\val _s2911_ : string -> maybe ((mword ty5 * mword ty5))\\ definition s2911 :: \ string \((5)Word.word*(5)Word.word)option \ where \ s2911 s29120 = ( (let s29130 = s29120 in if ((string_startswith s29130 (''sfence.vma''))) then (case ((string_drop s29130 ((string_length (''sfence.vma''))))) of s29140 => (case ((spc_matches_prefix0 s29140)) of Some ((_, s29150)) => (case ((string_drop s29140 s29150)) of s29160 => (case ((reg_name_matches_prefix s29160 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s29170)) => (case ((string_drop s29160 s29170)) of s29180 => (case ((sep_matches_prefix s29180)) of Some ((_, s29190)) => (case ((string_drop s29180 s29190)) of s29200 => (case ((reg_name_matches_prefix s29200 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s29210)) => (let p00 = (string_drop s29200 s29210) in if (((p00 = ('''')))) then Some (rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s29120 :: " string " \ \\val _s2899_ : string -> maybe ((mword ty4 * mword ty4))\\ definition s2899 :: \ string \((4)Word.word*(4)Word.word)option \ where \ s2899 s29000 = ( (let s29010 = s29000 in if ((string_startswith s29010 (''fence.tso''))) then (case ((string_drop s29010 ((string_length (''fence.tso''))))) of s29020 => (case ((spc_matches_prefix0 s29020)) of Some ((_, s29030)) => (case ((string_drop s29020 s29030)) of s29040 => (case ((fence_bits_matches_prefix s29040 :: (( 4 Word.word * ii)) option)) of Some ((pred, s29050)) => (case ((string_drop s29040 s29050)) of s29060 => (case ((sep_matches_prefix s29060)) of Some ((_, s29070)) => (case ((string_drop s29060 s29070)) of s29080 => (case ((fence_bits_matches_prefix s29080 :: (( 4 Word.word * ii)) option)) of Some ((succ, s29090)) => (let p00 = (string_drop s29080 s29090) in if (((p00 = ('''')))) then Some (pred, succ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s29000 :: " string " \ \\val _s2887_ : string -> maybe ((mword ty4 * mword ty4))\\ definition s2887 :: \ string \((4)Word.word*(4)Word.word)option \ where \ s2887 s28880 = ( (let s28890 = s28880 in if ((string_startswith s28890 (''fence''))) then (case ((string_drop s28890 ((string_length (''fence''))))) of s28900 => (case ((spc_matches_prefix0 s28900)) of Some ((_, s28910)) => (case ((string_drop s28900 s28910)) of s28920 => (case ((fence_bits_matches_prefix s28920 :: (( 4 Word.word * ii)) option)) of Some ((pred, s28930)) => (case ((string_drop s28920 s28930)) of s28940 => (case ((sep_matches_prefix s28940)) of Some ((_, s28950)) => (case ((string_drop s28940 s28950)) of s28960 => (case ((fence_bits_matches_prefix s28960 :: (( 4 Word.word * ii)) option)) of Some ((succ, s28970)) => (let p00 = (string_drop s28960 s28970) in if (((p00 = ('''')))) then Some (pred, succ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s28880 :: " string " \ \\val _s2870_ : string -> maybe ((sopw * mword ty5 * mword ty5 * mword ty5))\\ definition s2870 :: \ string \(sopw*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2870 s28720 = ( (case ((shiftiwop_mnemonic_matches_prefix s28720)) of Some ((op1, s28730)) => (case ((string_drop s28720 s28730)) of s28740 => (case ((spc_matches_prefix0 s28740)) of Some ((_, s28750)) => (case ((string_drop s28740 s28750)) of s28760 => (case ((reg_name_matches_prefix s28760 :: (( 5 Word.word * ii)) option)) of Some ((rd, s28770)) => (case ((string_drop s28760 s28770)) of s28780 => (case ((sep_matches_prefix s28780)) of Some ((_, s28790)) => (case ((string_drop s28780 s28790)) of s28800 => (case ((reg_name_matches_prefix s28800 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s28810)) => (case ((string_drop s28800 s28810)) of s28820 => (case ((sep_matches_prefix s28820)) of Some ((_, s28830)) => (case ((string_drop s28820 s28830)) of s28840 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s28840 :: (( 5 Word.word * ii)) option)) of Some ((shamt, s28850)) => (let p00 = (string_drop s28840 s28850) in if (((p00 = ('''')))) then Some (op1, rd, rs1, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s28720 :: " string " \ \\val _s2853_ : string -> maybe ((ropw * mword ty5 * mword ty5 * mword ty5))\\ definition s2853 :: \ string \(ropw*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2853 s28550 = ( (case ((rtypew_mnemonic_matches_prefix s28550)) of Some ((op1, s28560)) => (case ((string_drop s28550 s28560)) of s28570 => (case ((spc_matches_prefix0 s28570)) of Some ((_, s28580)) => (case ((string_drop s28570 s28580)) of s28590 => (case ((reg_name_matches_prefix s28590 :: (( 5 Word.word * ii)) option)) of Some ((rd, s28600)) => (case ((string_drop s28590 s28600)) of s28610 => (case ((sep_matches_prefix s28610)) of Some ((_, s28620)) => (case ((string_drop s28610 s28620)) of s28630 => (case ((reg_name_matches_prefix s28630 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s28640)) => (case ((string_drop s28630 s28640)) of s28650 => (case ((sep_matches_prefix s28650)) of Some ((_, s28660)) => (case ((string_drop s28650 s28660)) of s28670 => (case ((reg_name_matches_prefix s28670 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s28680)) => (let p00 = (string_drop s28670 s28680) in if (((p00 = ('''')))) then Some (op1, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s28550 :: " string " \ \\val _s2836_ : string -> maybe ((sop * mword ty5 * mword ty5 * mword ty5))\\ definition s2836 :: \ string \(sop*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2836 s28380 = ( (case ((shiftw_mnemonic_matches_prefix s28380)) of Some ((op1, s28390)) => (case ((string_drop s28380 s28390)) of s28400 => (case ((spc_matches_prefix0 s28400)) of Some ((_, s28410)) => (case ((string_drop s28400 s28410)) of s28420 => (case ((reg_name_matches_prefix s28420 :: (( 5 Word.word * ii)) option)) of Some ((rd, s28430)) => (case ((string_drop s28420 s28430)) of s28440 => (case ((sep_matches_prefix s28440)) of Some ((_, s28450)) => (case ((string_drop s28440 s28450)) of s28460 => (case ((reg_name_matches_prefix s28460 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s28470)) => (case ((string_drop s28460 s28470)) of s28480 => (case ((sep_matches_prefix s28480)) of Some ((_, s28490)) => (case ((string_drop s28480 s28490)) of s28500 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s28500 :: (( 5 Word.word * ii)) option)) of Some ((shamt, s28510)) => (let p00 = (string_drop s28500 s28510) in if (((p00 = ('''')))) then Some (op1, rd, rs1, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s28380 :: " string " \ \\val _s2820_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12))\\ definition s2820 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word)option \ where \ s2820 s28210 = ( (let s28220 = s28210 in if ((string_startswith s28220 (''addiw''))) then (case ((string_drop s28220 ((string_length (''addiw''))))) of s28230 => (case ((spc_matches_prefix0 s28230)) of Some ((_, s28240)) => (case ((string_drop s28230 s28240)) of s28250 => (case ((reg_name_matches_prefix s28250 :: (( 5 Word.word * ii)) option)) of Some ((rd, s28260)) => (case ((string_drop s28250 s28260)) of s28270 => (case ((sep_matches_prefix s28270)) of Some ((_, s28280)) => (case ((string_drop s28270 s28280)) of s28290 => (case ((reg_name_matches_prefix s28290 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s28300)) => (case ((string_drop s28290 s28300)) of s28310 => (case ((sep_matches_prefix s28310)) of Some ((_, s28320)) => (case ((string_drop s28310 s28320)) of s28330 => (case ((hex_bits_12_matches_prefix0 s28330 :: (( 12 Word.word * ii)) option)) of Some ((imm, s28340)) => (let p00 = (string_drop s28330 s28340) in if (((p00 = ('''')))) then Some (rd, rs1, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s28210 :: " string " \ \\val _s2792_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty12 * mword ty5))\\ definition s2792 :: \ string \(word_width*bool*bool*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s2792 s27930 = ( (let s27940 = s27930 in if ((string_startswith s27940 (''s''))) then (case ((string_drop s27940 ((string_length (''s''))))) of s27950 => (case ((size_mnemonic_matches_prefix s27950)) of Some ((size1, s27960)) => (case ((string_drop s27950 s27960)) of s27970 => (case ((maybe_aq_matches_prefix s27970)) of Some ((aq, s27980)) => (case ((string_drop s27970 s27980)) of s27990 => (case ((maybe_rl_matches_prefix s27990)) of Some ((rl, s28000)) => (case ((string_drop s27990 s28000)) of s28010 => (case ((spc_matches_prefix0 s28010)) of Some ((_, s28020)) => (case ((string_drop s28010 s28020)) of s28030 => (case ((reg_name_matches_prefix s28030 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s28040)) => (case ((string_drop s28030 s28040)) of s28050 => (case ((sep_matches_prefix s28050)) of Some ((_, s28060)) => (case ((string_drop s28050 s28060)) of s28070 => (case ((hex_bits_12_matches_prefix0 s28070 :: (( 12 Word.word * ii)) option)) of Some ((imm, s28080)) => (case ((string_drop s28070 s28080)) of s28090 => (case ((opt_spc_matches_prefix0 s28090)) of Some ((_, s28100)) => (let s28110 = (string_drop s28090 s28100) in if ((string_startswith s28110 (''(''))) then (case ((string_drop s28110 ((string_length (''(''))))) of s28120 => (case ((opt_spc_matches_prefix0 s28120)) of Some ((_, s28130)) => (case ((string_drop s28120 s28130)) of s28140 => (case ((reg_name_matches_prefix s28140 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s28150)) => (case ((string_drop s28140 s28150)) of s28160 => (case ((opt_spc_matches_prefix0 s28160)) of Some ((_, s28170)) => (let s28180 = (string_drop s28160 s28170) in if ((string_startswith s28180 ('')''))) then (let p00 = (string_drop s28180 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (size1, aq, rl, rs2, imm, rs1) else None) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s27930 :: " string " \ \\val _s2762_ : string -> maybe ((word_width * bool * bool * bool * mword ty5 * mword ty12 * mword ty5))\\ definition s2762 :: \ string \(word_width*bool*bool*bool*(5)Word.word*(12)Word.word*(5)Word.word)option \ where \ s2762 s27630 = ( (let s27640 = s27630 in if ((string_startswith s27640 (''l''))) then (case ((string_drop s27640 ((string_length (''l''))))) of s27650 => (case ((size_mnemonic_matches_prefix s27650)) of Some ((size1, s27660)) => (case ((string_drop s27650 s27660)) of s27670 => (case ((maybe_u_matches_prefix s27670)) of Some ((is_unsigned, s27680)) => (case ((string_drop s27670 s27680)) of s27690 => (case ((maybe_aq_matches_prefix s27690)) of Some ((aq, s27700)) => (case ((string_drop s27690 s27700)) of s27710 => (case ((maybe_rl_matches_prefix s27710)) of Some ((rl, s27720)) => (case ((string_drop s27710 s27720)) of s27730 => (case ((spc_matches_prefix0 s27730)) of Some ((_, s27740)) => (case ((string_drop s27730 s27740)) of s27750 => (case ((reg_name_matches_prefix s27750 :: (( 5 Word.word * ii)) option)) of Some ((rd, s27760)) => (case ((string_drop s27750 s27760)) of s27770 => (case ((sep_matches_prefix s27770)) of Some ((_, s27780)) => (case ((string_drop s27770 s27780)) of s27790 => (case ((hex_bits_12_matches_prefix0 s27790 :: (( 12 Word.word * ii)) option)) of Some ((imm, s27800)) => (case ((string_drop s27790 s27800)) of s27810 => (case ((opt_spc_matches_prefix0 s27810)) of Some ((_, s27820)) => (let s27830 = (string_drop s27810 s27820) in if ((string_startswith s27830 (''(''))) then (case ((string_drop s27830 ((string_length (''(''))))) of s27840 => (case ((opt_spc_matches_prefix0 s27840)) of Some ((_, s27850)) => (case ((string_drop s27840 s27850)) of s27860 => (case ((reg_name_matches_prefix s27860 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s27870)) => (case ((string_drop s27860 s27870)) of s27880 => (case ((opt_spc_matches_prefix0 s27880)) of Some ((_, s27890)) => (let s27900 = (string_drop s27880 s27890) in if ((string_startswith s27900 ('')''))) then (let p00 = (string_drop s27900 ((string_length ('')'')))) in if (((p00 = ('''')))) then Some (size1, is_unsigned, aq, rl, rd, imm, rs1) else None) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s27630 :: " string " \ \\val _s2745_ : string -> maybe ((rop * mword ty5 * mword ty5 * mword ty5))\\ definition s2745 :: \ string \(rop*(5)Word.word*(5)Word.word*(5)Word.word)option \ where \ s2745 s27470 = ( (case ((rtype_mnemonic_matches_prefix s27470)) of Some ((op1, s27480)) => (case ((string_drop s27470 s27480)) of s27490 => (case ((spc_matches_prefix0 s27490)) of Some ((_, s27500)) => (case ((string_drop s27490 s27500)) of s27510 => (case ((reg_name_matches_prefix s27510 :: (( 5 Word.word * ii)) option)) of Some ((rd, s27520)) => (case ((string_drop s27510 s27520)) of s27530 => (case ((sep_matches_prefix s27530)) of Some ((_, s27540)) => (case ((string_drop s27530 s27540)) of s27550 => (case ((reg_name_matches_prefix s27550 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s27560)) => (case ((string_drop s27550 s27560)) of s27570 => (case ((sep_matches_prefix s27570)) of Some ((_, s27580)) => (case ((string_drop s27570 s27580)) of s27590 => (case ((reg_name_matches_prefix s27590 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s27600)) => (let p00 = (string_drop s27590 s27600) in if (((p00 = ('''')))) then Some (op1, rd, rs1, rs2) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s27470 :: " string " \ \\val _s2728_ : string -> maybe ((sop * mword ty5 * mword ty5 * mword ty6))\\ definition s2728 :: \ string \(sop*(5)Word.word*(5)Word.word*(6)Word.word)option \ where \ s2728 s27300 = ( (case ((shiftiop_mnemonic_matches_prefix s27300)) of Some ((op1, s27310)) => (case ((string_drop s27300 s27310)) of s27320 => (case ((spc_matches_prefix0 s27320)) of Some ((_, s27330)) => (case ((string_drop s27320 s27330)) of s27340 => (case ((reg_name_matches_prefix s27340 :: (( 5 Word.word * ii)) option)) of Some ((rd, s27350)) => (case ((string_drop s27340 s27350)) of s27360 => (case ((sep_matches_prefix s27360)) of Some ((_, s27370)) => (case ((string_drop s27360 s27370)) of s27380 => (case ((reg_name_matches_prefix s27380 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s27390)) => (case ((string_drop s27380 s27390)) of s27400 => (case ((sep_matches_prefix s27400)) of Some ((_, s27410)) => (case ((string_drop s27400 s27410)) of s27420 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s27420 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s27430)) => (let p00 = (string_drop s27420 s27430) in if (((p00 = ('''')))) then Some (op1, rd, rs1, shamt) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s27300 :: " string " \ \\val _s2711_ : string -> maybe ((iop * mword ty5 * mword ty5 * mword ty12))\\ definition s2711 :: \ string \(iop*(5)Word.word*(5)Word.word*(12)Word.word)option \ where \ s2711 s27130 = ( (case ((itype_mnemonic_matches_prefix s27130)) of Some ((op1, s27140)) => (case ((string_drop s27130 s27140)) of s27150 => (case ((spc_matches_prefix0 s27150)) of Some ((_, s27160)) => (case ((string_drop s27150 s27160)) of s27170 => (case ((reg_name_matches_prefix s27170 :: (( 5 Word.word * ii)) option)) of Some ((rd, s27180)) => (case ((string_drop s27170 s27180)) of s27190 => (case ((sep_matches_prefix s27190)) of Some ((_, s27200)) => (case ((string_drop s27190 s27200)) of s27210 => (case ((reg_name_matches_prefix s27210 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s27220)) => (case ((string_drop s27210 s27220)) of s27230 => (case ((sep_matches_prefix s27230)) of Some ((_, s27240)) => (case ((string_drop s27230 s27240)) of s27250 => (case ((hex_bits_12_matches_prefix0 s27250 :: (( 12 Word.word * ii)) option)) of Some ((imm, s27260)) => (let p00 = (string_drop s27250 s27260) in if (((p00 = ('''')))) then Some (op1, rd, rs1, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s27130 :: " string " \ \\val _s2694_ : string -> maybe ((bop * mword ty5 * mword ty5 * mword ty13))\\ definition s2694 :: \ string \(bop*(5)Word.word*(5)Word.word*(13)Word.word)option \ where \ s2694 s26960 = ( (case ((btype_mnemonic_matches_prefix s26960)) of Some ((op1, s26970)) => (case ((string_drop s26960 s26970)) of s26980 => (case ((spc_matches_prefix0 s26980)) of Some ((_, s26990)) => (case ((string_drop s26980 s26990)) of s27000 => (case ((reg_name_matches_prefix s27000 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s27010)) => (case ((string_drop s27000 s27010)) of s27020 => (case ((sep_matches_prefix s27020)) of Some ((_, s27030)) => (case ((string_drop s27020 s27030)) of s27040 => (case ((reg_name_matches_prefix s27040 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s27050)) => (case ((string_drop s27040 s27050)) of s27060 => (case ((sep_matches_prefix s27060)) of Some ((_, s27070)) => (case ((string_drop s27060 s27070)) of s27080 => (case ((hex_bits_13_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s27080 :: (( 13 Word.word * ii)) option)) of Some ((imm, s27090)) => (let p00 = (string_drop s27080 s27090) in if (((p00 = ('''')))) then Some (op1, rs1, rs2, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s26960 :: " string " \ \\val _s2678_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12))\\ definition s2678 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word)option \ where \ s2678 s26790 = ( (let s26800 = s26790 in if ((string_startswith s26800 (''jalr''))) then (case ((string_drop s26800 ((string_length (''jalr''))))) of s26810 => (case ((spc_matches_prefix0 s26810)) of Some ((_, s26820)) => (case ((string_drop s26810 s26820)) of s26830 => (case ((reg_name_matches_prefix s26830 :: (( 5 Word.word * ii)) option)) of Some ((rd, s26840)) => (case ((string_drop s26830 s26840)) of s26850 => (case ((sep_matches_prefix s26850)) of Some ((_, s26860)) => (case ((string_drop s26850 s26860)) of s26870 => (case ((reg_name_matches_prefix s26870 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s26880)) => (case ((string_drop s26870 s26880)) of s26890 => (case ((sep_matches_prefix s26890)) of Some ((_, s26900)) => (case ((string_drop s26890 s26900)) of s26910 => (case ((hex_bits_12_matches_prefix0 s26910 :: (( 12 Word.word * ii)) option)) of Some ((imm, s26920)) => (let p00 = (string_drop s26910 s26920) in if (((p00 = ('''')))) then Some (rd, rs1, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s26790 :: " string " \ \\val _s2666_ : string -> maybe ((mword ty5 * mword ty21))\\ definition s2666 :: \ string \((5)Word.word*(21)Word.word)option \ where \ s2666 s26670 = ( (let s26680 = s26670 in if ((string_startswith s26680 (''jal''))) then (case ((string_drop s26680 ((string_length (''jal''))))) of s26690 => (case ((spc_matches_prefix0 s26690)) of Some ((_, s26700)) => (case ((string_drop s26690 s26700)) of s26710 => (case ((reg_name_matches_prefix s26710 :: (( 5 Word.word * ii)) option)) of Some ((rd, s26720)) => (case ((string_drop s26710 s26720)) of s26730 => (case ((sep_matches_prefix s26730)) of Some ((_, s26740)) => (case ((string_drop s26730 s26740)) of s26750 => (case ((hex_bits_21_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s26750 :: (( 21 Word.word * ii)) option)) of Some ((imm, s26760)) => (let p00 = (string_drop s26750 s26760) in if (((p00 = ('''')))) then Some (rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s26670 :: " string " \ \\val _s2653_ : string -> maybe ((uop * mword ty5 * mword ty20))\\ definition s2653 :: \ string \(uop*(5)Word.word*(20)Word.word)option \ where \ s2653 s26550 = ( (case ((utype_mnemonic_matches_prefix s26550)) of Some ((op1, s26560)) => (case ((string_drop s26550 s26560)) of s26570 => (case ((spc_matches_prefix0 s26570)) of Some ((_, s26580)) => (case ((string_drop s26570 s26580)) of s26590 => (case ((reg_name_matches_prefix s26590 :: (( 5 Word.word * ii)) option)) of Some ((rd, s26600)) => (case ((string_drop s26590 s26600)) of s26610 => (case ((sep_matches_prefix s26610)) of Some ((_, s26620)) => (case ((string_drop s26610 s26620)) of s26630 => (case ((hex_bits_20_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s26630 :: (( 20 Word.word * ii)) option)) of Some ((imm, s26640)) => (let p00 = (string_drop s26630 s26640) in if (((p00 = ('''')))) then Some (op1, rd, imm) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s26550 :: " string " definition assembly_backwards_matches :: \ string \ bool \ where \ assembly_backwards_matches arg1 = ( (let s26650 = arg1 in if ((case ((s2653 s26650 :: ((uop * 5 Word.word * 20 Word.word))option)) of Some ((op1, rd, imm)) => True | _ => False )) then (case (s2653 s26650 :: (( uop * 5 Word.word * 20 Word.word)) option) of (Some ((op1, rd, imm))) => True ) else if ((case ((s2666 s26650 :: (( 5 Word.word * 21 Word.word))option)) of Some ((rd, imm)) => True | _ => False )) then (case (s2666 s26650 :: (( 5 Word.word * 21 Word.word)) option) of (Some ((rd, imm))) => True ) else if ((case ((s2678 s26650 :: (( 5 Word.word * 5 Word.word * 12 Word.word))option)) of Some ((rd, rs1, imm)) => True | _ => False )) then (case (s2678 s26650 :: (( 5 Word.word * 5 Word.word * 12 Word.word)) option) of (Some ((rd, rs1, imm))) => True ) else if ((case ((s2694 s26650 :: ((bop * 5 Word.word * 5 Word.word * 13 Word.word))option)) of Some ((op1, rs1, rs2, imm)) => True | _ => False )) then (case (s2694 s26650 :: (( bop * 5 Word.word * 5 Word.word * 13 Word.word)) option) of (Some ((op1, rs1, rs2, imm))) => True ) else if ((case ((s2711 s26650 :: ((iop * 5 Word.word * 5 Word.word * 12 Word.word))option)) of Some ((op1, rd, rs1, imm)) => True | _ => False )) then (case (s2711 s26650 :: (( iop * 5 Word.word * 5 Word.word * 12 Word.word)) option) of (Some ((op1, rd, rs1, imm))) => True ) else if ((case ((s2728 s26650 :: ((sop * 5 Word.word * 5 Word.word * 6 Word.word))option)) of Some ((op1, rd, rs1, shamt)) => True | _ => False )) then (case (s2728 s26650 :: (( sop * 5 Word.word * 5 Word.word * 6 Word.word)) option) of (Some ((op1, rd, rs1, shamt))) => True ) else if ((case ((s2745 s26650 :: ((rop * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, rd, rs1, rs2)) => True | _ => False )) then (case (s2745 s26650 :: (( rop * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, rd, rs1, rs2))) => True ) else if ((case ((s2762 s26650 :: ((word_width * bool * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((size1, is_unsigned, aq, rl, rd, imm, rs1)) => True | _ => False )) then (case (s2762 s26650 :: (( word_width * bool * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((size1, is_unsigned, aq, rl, rd, imm, rs1))) => True ) else if ((case ((s2792 s26650 :: ((word_width * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((size1, aq, rl, rs2, imm, rs1)) => True | _ => False )) then (case (s2792 s26650 :: (( word_width * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((size1, aq, rl, rs2, imm, rs1))) => True ) else if ((case ((s2820 s26650 :: (( 5 Word.word * 5 Word.word * 12 Word.word))option)) of Some ((rd, rs1, imm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s2820 s26650 :: (( 5 Word.word * 5 Word.word * 12 Word.word)) option) of (Some ((rd, rs1, imm))) => True ) else if ((case ((s2836 s26650 :: ((sop * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, rd, rs1, shamt)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s2836 s26650 :: (( sop * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, rd, rs1, shamt))) => True ) else if ((case ((s2853 s26650 :: ((ropw * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, rd, rs1, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s2853 s26650 :: (( ropw * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, rd, rs1, rs2))) => True ) else if ((case ((s2870 s26650 :: ((sopw * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, rd, rs1, shamt)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s2870 s26650 :: (( sopw * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, rd, rs1, shamt))) => True ) else if ((case ((s2887 s26650 :: (( 4 Word.word * 4 Word.word))option)) of Some ((pred, succ)) => True | _ => False )) then (case (s2887 s26650 :: (( 4 Word.word * 4 Word.word)) option) of (Some ((pred, succ))) => True ) else if ((case ((s2899 s26650 :: (( 4 Word.word * 4 Word.word))option)) of Some ((pred, succ)) => True | _ => False )) then (case (s2899 s26650 :: (( 4 Word.word * 4 Word.word)) option) of (Some ((pred, succ))) => True ) else if (((s26650 = (''fence.i'')))) then True else if (((s26650 = (''ecall'')))) then True else if (((s26650 = (''mret'')))) then True else if (((s26650 = (''sret'')))) then True else if (((s26650 = (''ebreak'')))) then True else if (((s26650 = (''wfi'')))) then True else if ((case ((s2911 s26650 :: (( 5 Word.word * 5 Word.word))option)) of Some ((rs1, rs2)) => True | _ => False )) then (case (s2911 s26650 :: (( 5 Word.word * 5 Word.word)) option) of (Some ((rs1, rs2))) => True ) else if ((case ((s2923 s26650 :: ((word_width * bool * bool * 5 Word.word * 5 Word.word))option)) of Some ((size1, aq, rl, rd, rs1)) => True | _ => False )) then (case (s2923 s26650 :: (( word_width * bool * bool * 5 Word.word * 5 Word.word)) option) of (Some ((size1, aq, rl, rd, rs1))) => True ) else if ((case ((s2941 s26650 :: ((word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((size1, aq, rl, rd, rs1, rs2)) => True | _ => False )) then (case (s2941 s26650 :: (( word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((size1, aq, rl, rd, rs1, rs2))) => True ) else if ((case ((s2963 s26650 :: ((amoop * word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((op1, width, aq, rl, rd, rs2, rs1)) => True | _ => False )) then (case (s2963 s26650 :: (( amoop * word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((op1, width, aq, rl, rd, rs2, rs1))) => True ) else if (((s26650 = (''c.nop'')))) then True else if ((case ((s2989 s26650 :: (( 3 Word.word * 8 Word.word))option)) of Some ((rdc, nzimm)) => (nzimm \ ( 0x00 :: 8 Word.word)) | _ => False )) then (case (s2989 s26650 :: (( 3 Word.word * 8 Word.word)) option) of (Some ((rdc, nzimm))) => True ) else if ((case ((s3001 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rdc, rsc, uimm)) => True | _ => False )) then (case (s3001 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rdc, rsc, uimm))) => True ) else if ((case ((s3017 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rdc, rsc, uimm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3017 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rdc, rsc, uimm))) => True ) else if ((case ((s3033 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rsc1, rsc2, uimm)) => True | _ => False )) then (case (s3033 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rsc1, rsc2, uimm))) => True ) else if ((case ((s3049 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rsc1, rsc2, uimm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3049 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rsc1, rsc2, uimm))) => True ) else if ((case ((s3065 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rsd, nzi)) => ((((nzi \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))) | _ => False )) then (case (s3065 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rsd, nzi))) => True ) else if ((case ((s3077 s26650 :: ( 11 Word.word)option)) of Some (imm) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s3077 s26650 :: ( 11 Word.word) option) of (Some (imm)) => True ) else if ((case ((s3085 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rsd, imm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3085 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rsd, imm))) => True ) else if ((case ((s3097 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, imm)) => (rd \ zreg) | _ => False )) then (case (s3097 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, imm))) => True ) else if ((case ((s3109 s26650 :: ( 6 Word.word)option)) of Some (imm) => (imm \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s3109 s26650 :: ( 6 Word.word) option) of (Some (imm)) => True ) else if ((case ((s3117 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, imm)) => ((((rd \ zreg))) \ ((((((rd \ sp))) \ (((imm \ ( 0b000000 :: 6 Word.word)))))))) | _ => False )) then (case (s3117 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, imm))) => True ) else if ((case ((s3129 s26650 :: (( 3 Word.word * 6 Word.word))option)) of Some ((rsd, shamt)) => (shamt \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s3129 s26650 :: (( 3 Word.word * 6 Word.word)) option) of (Some ((rsd, shamt))) => True ) else if ((case ((s3141 s26650 :: (( 3 Word.word * 6 Word.word))option)) of Some ((rsd, shamt)) => (shamt \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s3141 s26650 :: (( 3 Word.word * 6 Word.word)) option) of (Some ((rsd, shamt))) => True ) else if ((case ((s3153 s26650 :: (( 3 Word.word * 6 Word.word))option)) of Some ((rsd, imm)) => True | _ => False )) then (case (s3153 s26650 :: (( 3 Word.word * 6 Word.word)) option) of (Some ((rsd, imm))) => True ) else if ((case ((s3165 s26650 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => True | _ => False )) then (case (s3165 s26650 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => True ) else if ((case ((s3177 s26650 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => True | _ => False )) then (case (s3177 s26650 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => True ) else if ((case ((s3189 s26650 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => True | _ => False )) then (case (s3189 s26650 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => True ) else if ((case ((s3201 s26650 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => True | _ => False )) then (case (s3201 s26650 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => True ) else if ((case ((s3213 s26650 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3213 s26650 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => True ) else if ((case ((s3225 s26650 :: (( 3 Word.word * 3 Word.word))option)) of Some ((rsd, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3225 s26650 :: (( 3 Word.word * 3 Word.word)) option) of (Some ((rsd, rs2))) => True ) else if ((case ((s3237 s26650 :: ( 11 Word.word)option)) of Some (imm) => True | _ => False )) then (case (s3237 s26650 :: ( 11 Word.word) option) of (Some (imm)) => True ) else if ((case ((s3245 s26650 :: (( 3 Word.word * 8 Word.word))option)) of Some ((rs, imm)) => True | _ => False )) then (case (s3245 s26650 :: (( 3 Word.word * 8 Word.word)) option) of (Some ((rs, imm))) => True ) else if ((case ((s3257 s26650 :: (( 3 Word.word * 8 Word.word))option)) of Some ((rs, imm)) => True | _ => False )) then (case (s3257 s26650 :: (( 3 Word.word * 8 Word.word)) option) of (Some ((rs, imm))) => True ) else if ((case ((s3269 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rsd, shamt)) => ((((shamt \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))) | _ => False )) then (case (s3269 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rsd, shamt))) => True ) else if ((case ((s3281 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, uimm)) => (rd \ zreg) | _ => False )) then (case (s3281 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, uimm))) => True ) else if ((case ((s3293 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, uimm)) => ((((rd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))) | _ => False )) then (case (s3293 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, uimm))) => True ) else if ((case ((s3305 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, uimm)) => True | _ => False )) then (case (s3305 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, uimm))) => True ) else if ((case ((s3317 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rs2, uimm)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3317 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rs2, uimm))) => True ) else if ((case ((s3329 s26650 :: ( 5 Word.word)option)) of Some (rs1) => (rs1 \ zreg) | _ => False )) then (case (s3329 s26650 :: ( 5 Word.word) option) of (Some (rs1)) => True ) else if ((case ((s3337 s26650 :: ( 5 Word.word)option)) of Some (rs1) => (rs1 \ zreg) | _ => False )) then (case (s3337 s26650 :: ( 5 Word.word) option) of (Some (rs1)) => True ) else if ((case ((s3345 s26650 :: (( 5 Word.word * 5 Word.word))option)) of Some ((rd, rs2)) => ((((rd \ zreg))) \ (((rs2 \ zreg)))) | _ => False )) then (case (s3345 s26650 :: (( 5 Word.word * 5 Word.word)) option) of (Some ((rd, rs2))) => True ) else if (((s26650 = (''c.ebreak'')))) then True else if ((case ((s3357 s26650 :: (( 5 Word.word * 5 Word.word))option)) of Some ((rsd, rs2)) => ((((rsd \ zreg))) \ (((rs2 \ zreg)))) | _ => False )) then (case (s3357 s26650 :: (( 5 Word.word * 5 Word.word)) option) of (Some ((rsd, rs2))) => True ) else if ((case ((s3369 s26650 :: ((bool * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((high, signed1, signed2, rd, rs1, rs2)) => True | _ => False )) then (case (s3369 s26650 :: (( bool * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((high, signed1, signed2, rd, rs1, rs2))) => True ) else if ((case ((s3386 s26650 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((s, rd, rs1, rs2)) => True | _ => False )) then (case (s3386 s26650 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((s, rd, rs1, rs2))) => True ) else if ((case ((s3404 s26650 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((s, rd, rs1, rs2)) => True | _ => False )) then (case (s3404 s26650 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((s, rd, rs1, rs2))) => True ) else if ((case ((s3422 s26650 :: (( 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((rd, rs1, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3422 s26650 :: (( 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((rd, rs1, rs2))) => True ) else if ((case ((s3438 s26650 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((s, rd, rs1, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3438 s26650 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((s, rd, rs1, rs2))) => True ) else if ((case ((s3457 s26650 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((s, rd, rs1, rs2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s3457 s26650 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((s, rd, rs1, rs2))) => True ) else if ((case ((s3476 s26650 :: ((csrop * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((op1, rd, csr, rs1)) => True | _ => False )) then (case (s3476 s26650 :: (( csrop * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((op1, rd, csr, rs1))) => True ) else if ((case ((s3494 s26650 :: ((csrop * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((op1, rd, csr, rs1)) => True | _ => False )) then (case (s3494 s26650 :: (( csrop * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((op1, rd, csr, rs1))) => True ) else if (((s26650 = (''uret'')))) then True else if ((case ((s3511 s26650 :: ( 6 Word.word)option)) of Some (imm) => True | _ => False )) then (case (s3511 s26650 :: ( 6 Word.word) option) of (Some (imm)) => True ) else if ((case ((s3517 s26650 :: ( 5 Word.word)option)) of Some (rsd) => (rsd \ zreg) | _ => False )) then (case (s3517 s26650 :: ( 5 Word.word) option) of (Some (rsd)) => True ) else if ((case ((s3523 s26650 :: ( 6 Word.word)option)) of Some (imm) => True | _ => False )) then (case (s3523 s26650 :: ( 6 Word.word) option) of (Some (imm)) => True ) else if ((case ((s3529 s26650 :: ( 6 Word.word)option)) of Some (imm) => (imm \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s3529 s26650 :: ( 6 Word.word) option) of (Some (imm)) => True ) else if ((case ((s3535 s26650 :: ( 5 Word.word)option)) of Some (rs2) => (rs2 \ zreg) | _ => False )) then (case (s3535 s26650 :: ( 5 Word.word) option) of (Some (rs2)) => True ) else if ((case ((s3541 s26650 :: ( 5 Word.word)option)) of Some (rs2) => (rs2 \ zreg) | _ => False )) then (case (s3541 s26650 :: ( 5 Word.word) option) of (Some (rs2)) => True ) else if ((case ((s3547 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rsd, shamt)) => ((((shamt = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))) | _ => False )) then (case (s3547 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rsd, shamt))) => True ) else if ((case ((s3556 s26650 :: ( 3 Word.word)option)) of Some (rsd) => True | _ => False )) then (case (s3556 s26650 :: ( 3 Word.word) option) of (Some (rsd)) => True ) else if ((case ((s3562 s26650 :: ( 3 Word.word)option)) of Some (rsd) => True | _ => False )) then (case (s3562 s26650 :: ( 3 Word.word) option) of (Some (rsd)) => True ) else if ((case ((s3568 s26650 :: (( 4 Word.word * 4 Word.word * 5 Word.word * 5 Word.word * 4 Word.word))option)) of Some ((pred, succ, rs, rd, fm)) => (((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word)))))))) | _ => False )) then (case (s3568 s26650 :: (( 4 Word.word * 4 Word.word * 5 Word.word * 5 Word.word * 4 Word.word)) option) of (Some ((pred, succ, rs, rd, fm))) => True ) else if ((case ((s3586 s26650 :: (( 5 Word.word * 5 Word.word * 12 Word.word))option)) of Some ((rd, rs, imm)) => ((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg))))))) | _ => False )) then (case (s3586 s26650 :: (( 5 Word.word * 5 Word.word * 12 Word.word)) option) of (Some ((rd, rs, imm))) => True ) else if ((case ((s3598 s26650 :: ((word_width * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((width, rd, imm, rs1)) => True | _ => False )) then (case (s3598 s26650 :: (( word_width * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((width, rd, imm, rs1))) => True ) else if ((case ((s3622 s26650 :: ((word_width * 5 Word.word * 12 Word.word * 5 Word.word))option)) of Some ((width, rs2, imm, rs1)) => True | _ => False )) then (case (s3622 s26650 :: (( word_width * 5 Word.word * 12 Word.word * 5 Word.word)) option) of (Some ((width, rs2, imm, rs1))) => True ) else if ((case ((s3646 s26650 :: ((f_madd_op_S * 5 Word.word * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((op1, rd, rs1, rs2, rs3, rm)) => True | _ => False )) then (case (s3646 s26650 :: (( f_madd_op_S * 5 Word.word * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((op1, rd, rs1, rs2, rs3, rm))) => True ) else if ((case ((s3671 s26650 :: ((f_bin_rm_op_S * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((op1, rd, rs1, rs2, rm)) => True | _ => False )) then (case (s3671 s26650 :: (( f_bin_rm_op_S * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((op1, rd, rs1, rs2, rm))) => True ) else if ((case ((s3692 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FSQRT_S, rd, rs1, rm)) => True | _ => False )) then (case (s3692 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FSQRT_S, rd, rs1, rm))) => True ) else if ((case ((s3709 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_W_S, rd, rs1, rm)) => True | _ => False )) then (case (s3709 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_W_S, rd, rs1, rm))) => True ) else if ((case ((s3726 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_WU_S, rd, rs1, rm)) => True | _ => False )) then (case (s3726 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_WU_S, rd, rs1, rm))) => True ) else if ((case ((s3743 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_S_W, rd, rs1, rm)) => True | _ => False )) then (case (s3743 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_S_W, rd, rs1, rm))) => True ) else if ((case ((s3760 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_S_WU, rd, rs1, rm)) => True | _ => False )) then (case (s3760 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_S_WU, rd, rs1, rm))) => True ) else if ((case ((s3777 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_L_S, rd, rs1, rm)) => True | _ => False )) then (case (s3777 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_L_S, rd, rs1, rm))) => True ) else if ((case ((s3794 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_LU_S, rd, rs1, rm)) => True | _ => False )) then (case (s3794 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_LU_S, rd, rs1, rm))) => True ) else if ((case ((s3811 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_S_L, rd, rs1, rm)) => True | _ => False )) then (case (s3811 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_S_L, rd, rs1, rm))) => True ) else if ((case ((s3828 s26650 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode))option)) of Some ((FCVT_S_LU, rd, rs1, rm)) => True | _ => False )) then (case (s3828 s26650 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode)) option) of (Some ((FCVT_S_LU, rd, rs1, rm))) => True ) else if ((case ((s3845 s26650 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FSGNJ_S, rd, rs1, rs2)) => True | _ => False )) then (case (s3845 s26650 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FSGNJ_S, rd, rs1, rs2))) => True ) else if ((case ((s3862 s26650 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FSGNJN_S, rd, rs1, rs2)) => True | _ => False )) then (case (s3862 s26650 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FSGNJN_S, rd, rs1, rs2))) => True ) else if ((case ((s3879 s26650 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FSGNJX_S, rd, rs1, rs2)) => True | _ => False )) then (case (s3879 s26650 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FSGNJX_S, rd, rs1, rs2))) => True ) else if ((case ((s3896 s26650 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FMIN_S, rd, rs1, rs2)) => True | _ => False )) then (case (s3896 s26650 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FMIN_S, rd, rs1, rs2))) => True ) else if ((case ((s3913 s26650 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FMAX_S, rd, rs1, rs2)) => True | _ => False )) then (case (s3913 s26650 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FMAX_S, rd, rs1, rs2))) => True ) else if ((case ((s3930 s26650 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FEQ_S, rd, rs1, rs2)) => True | _ => False )) then (case (s3930 s26650 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FEQ_S, rd, rs1, rs2))) => True ) else if ((case ((s3947 s26650 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FLT_S, rd, rs1, rs2)) => True | _ => False )) then (case (s3947 s26650 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FLT_S, rd, rs1, rs2))) => True ) else if ((case ((s3964 s26650 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word))option)) of Some ((FLE_S, rd, rs1, rs2)) => True | _ => False )) then (case (s3964 s26650 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word)) option) of (Some ((FLE_S, rd, rs1, rs2))) => True ) else if ((case ((s3981 s26650 :: ((f_un_op_S * 5 Word.word * 5 Word.word))option)) of Some ((FMV_X_W, rd, rs1)) => True | _ => False )) then (case (s3981 s26650 :: (( f_un_op_S * 5 Word.word * 5 Word.word)) option) of (Some ((FMV_X_W, rd, rs1))) => True ) else if ((case ((s3994 s26650 :: ((f_un_op_S * 5 Word.word * 5 Word.word))option)) of Some ((FMV_W_X, rd, rs1)) => True | _ => False )) then (case (s3994 s26650 :: (( f_un_op_S * 5 Word.word * 5 Word.word)) option) of (Some ((FMV_W_X, rd, rs1))) => True ) else if ((case ((s4007 s26650 :: ((f_un_op_S * 5 Word.word * 5 Word.word))option)) of Some ((FCLASS_S, rd, rs1)) => True | _ => False )) then (case (s4007 s26650 :: (( f_un_op_S * 5 Word.word * 5 Word.word)) option) of (Some ((FCLASS_S, rd, rs1))) => True ) else if ((case ((s4020 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, imm)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s4020 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, imm))) => True ) else if ((case ((s4032 s26650 :: (( 5 Word.word * 6 Word.word))option)) of Some ((rd, uimm)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s4032 s26650 :: (( 5 Word.word * 6 Word.word)) option) of (Some ((rd, uimm))) => True ) else if ((case ((s4044 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rdc, rsc, uimm)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s4044 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rdc, rsc, uimm))) => True ) else if ((case ((s4060 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word))option)) of Some ((rsc1, rsc2, uimm)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s4060 s26650 :: (( 3 Word.word * 3 Word.word * 5 Word.word)) option) of (Some ((rsc1, rsc2, uimm))) => True ) else if ((case ((s4076 s26650 :: ( 32 Word.word)option)) of Some (s) => True | _ => False )) then (case (s4076 s26650 :: ( 32 Word.word) option) of (Some (s)) => True ) else if ((case ((s4084 s26650 :: ( 16 Word.word)option)) of Some (s) => True | _ => False )) then (case (s4084 s26650 :: ( 16 Word.word) option) of (Some (s)) => True ) else False))\ for arg1 :: " string " \ \\val _s5559_ : string -> maybe ((mword ty16 * string))\\ definition s5559 :: \ string \((16)Word.word*string)option \ where \ s5559 s55600 = ( (let s55610 = s55600 in if ((string_startswith s55610 (''c.illegal''))) then (case ((string_drop s55610 ((string_length (''c.illegal''))))) of s55620 => (case ((spc_matches_prefix0 s55620)) of Some ((_, s55630)) => (case ((string_drop s55620 s55630)) of s55640 => (case ((hex_bits_16_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s55640 :: (( 16 Word.word * ii)) option)) of Some ((s, s55650)) => (case ((string_drop s55640 s55650)) of s2 => Some (s, s2) ) | _ => None ) ) | _ => None ) ) else None))\ for s55600 :: " string " \ \\val _s5551_ : string -> maybe ((mword ty32 * string))\\ definition s5551 :: \ string \((32)Word.word*string)option \ where \ s5551 s55520 = ( (let s55530 = s55520 in if ((string_startswith s55530 (''illegal''))) then (case ((string_drop s55530 ((string_length (''illegal''))))) of s55540 => (case ((spc_matches_prefix0 s55540)) of Some ((_, s55550)) => (case ((string_drop s55540 s55550)) of s55560 => (case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s55560 :: (( 32 Word.word * ii)) option)) of Some ((s, s55570)) => (case ((string_drop s55560 s55570)) of s2 => Some (s, s2) ) | _ => None ) ) | _ => None ) ) else None))\ for s55520 :: " string " \ \\val _s5535_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5 * string))\\ definition s5535 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word*string)option \ where \ s5535 s55360 = ( (let s55370 = s55360 in if ((string_startswith s55370 (''c.fsw''))) then (case ((string_drop s55370 ((string_length (''c.fsw''))))) of s55380 => (case ((spc_matches_prefix0 s55380)) of Some ((_, s55390)) => (case ((string_drop s55380 s55390)) of s55400 => (case ((creg_name_matches_prefix s55400 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s55410)) => (case ((string_drop s55400 s55410)) of s55420 => (case ((sep_matches_prefix s55420)) of Some ((_, s55430)) => (case ((string_drop s55420 s55430)) of s55440 => (case ((creg_name_matches_prefix s55440 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s55450)) => (case ((string_drop s55440 s55450)) of s55460 => (case ((sep_matches_prefix s55460)) of Some ((_, s55470)) => (case ((string_drop s55460 s55470)) of s55480 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s55480 :: (( 7 Word.word * ii)) option)) of Some ((v__1212, s55490)) => if (((((subrange_vec_dec v__1212 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1212 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1212 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (case ((string_drop s55480 s55490)) of s1 => Some (rsc1, rsc2, uimm, s1) ))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s55360 :: " string " \ \\val _s5519_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5 * string))\\ definition s5519 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word*string)option \ where \ s5519 s55200 = ( (let s55210 = s55200 in if ((string_startswith s55210 (''c.flw''))) then (case ((string_drop s55210 ((string_length (''c.flw''))))) of s55220 => (case ((spc_matches_prefix0 s55220)) of Some ((_, s55230)) => (case ((string_drop s55220 s55230)) of s55240 => (case ((creg_name_matches_prefix s55240 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s55250)) => (case ((string_drop s55240 s55250)) of s55260 => (case ((sep_matches_prefix s55260)) of Some ((_, s55270)) => (case ((string_drop s55260 s55270)) of s55280 => (case ((creg_name_matches_prefix s55280 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s55290)) => (case ((string_drop s55280 s55290)) of s55300 => (case ((sep_matches_prefix s55300)) of Some ((_, s55310)) => (case ((string_drop s55300 s55310)) of s55320 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s55320 :: (( 7 Word.word * ii)) option)) of Some ((v__1214, s55330)) => if (((((subrange_vec_dec v__1214 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1214 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1214 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (case ((string_drop s55320 s55330)) of s1 => Some (rdc, rsc, uimm, s1) ))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s55200 :: " string " \ \\val _s5507_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s5507 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s5507 s55080 = ( (let s55090 = s55080 in if ((string_startswith s55090 (''c.fswsp''))) then (case ((string_drop s55090 ((string_length (''c.fswsp''))))) of s55100 => (case ((spc_matches_prefix0 s55100)) of Some ((_, s55110)) => (case ((string_drop s55100 s55110)) of s55120 => (case ((reg_name_matches_prefix s55120 :: (( 5 Word.word * ii)) option)) of Some ((rd, s55130)) => (case ((string_drop s55120 s55130)) of s55140 => (case ((sep_matches_prefix s55140)) of Some ((_, s55150)) => (case ((string_drop s55140 s55150)) of s55160 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s55160 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s55170)) => (case ((string_drop s55160 s55170)) of s1 => Some (rd, uimm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s55080 :: " string " \ \\val _s5495_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s5495 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s5495 s54960 = ( (let s54970 = s54960 in if ((string_startswith s54970 (''c.flwsp''))) then (case ((string_drop s54970 ((string_length (''c.flwsp''))))) of s54980 => (case ((spc_matches_prefix0 s54980)) of Some ((_, s54990)) => (case ((string_drop s54980 s54990)) of s55000 => (case ((reg_name_matches_prefix s55000 :: (( 5 Word.word * ii)) option)) of Some ((rd, s55010)) => (case ((string_drop s55000 s55010)) of s55020 => (case ((sep_matches_prefix s55020)) of Some ((_, s55030)) => (case ((string_drop s55020 s55030)) of s55040 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s55040 :: (( 6 Word.word * ii)) option)) of Some ((imm, s55050)) => (case ((string_drop s55040 s55050)) of s1 => Some (rd, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s54960 :: " string " \ \\val _s5482_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5 * string))\\ definition s5482 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word*string)option \ where \ s5482 s54840 = ( (case ((f_un_type_mnemonic_S_matches_prefix s54840)) of Some ((FCLASS_S, s54850)) => (case ((string_drop s54840 s54850)) of s54860 => (case ((spc_matches_prefix0 s54860)) of Some ((_, s54870)) => (case ((string_drop s54860 s54870)) of s54880 => (case ((reg_name_matches_prefix s54880 :: (( 5 Word.word * ii)) option)) of Some ((rd, s54890)) => (case ((string_drop s54880 s54890)) of s54900 => (case ((sep_matches_prefix s54900)) of Some ((_, s54910)) => (case ((string_drop s54900 s54910)) of s54920 => (case ((freg_name_matches_prefix s54920 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s54930)) => (case ((string_drop s54920 s54930)) of s1 => Some (FCLASS_S, rd, rs1, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s54840 :: " string " \ \\val _s5469_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5 * string))\\ definition s5469 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word*string)option \ where \ s5469 s54710 = ( (case ((f_un_type_mnemonic_S_matches_prefix s54710)) of Some ((FMV_W_X, s54720)) => (case ((string_drop s54710 s54720)) of s54730 => (case ((spc_matches_prefix0 s54730)) of Some ((_, s54740)) => (case ((string_drop s54730 s54740)) of s54750 => (case ((freg_name_matches_prefix s54750 :: (( 5 Word.word * ii)) option)) of Some ((rd, s54760)) => (case ((string_drop s54750 s54760)) of s54770 => (case ((sep_matches_prefix s54770)) of Some ((_, s54780)) => (case ((string_drop s54770 s54780)) of s54790 => (case ((reg_name_matches_prefix s54790 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s54800)) => (case ((string_drop s54790 s54800)) of s1 => Some (FMV_W_X, rd, rs1, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s54710 :: " string " \ \\val _s5456_ : string -> maybe ((f_un_op_S * mword ty5 * mword ty5 * string))\\ definition s5456 :: \ string \(f_un_op_S*(5)Word.word*(5)Word.word*string)option \ where \ s5456 s54580 = ( (case ((f_un_type_mnemonic_S_matches_prefix s54580)) of Some ((FMV_X_W, s54590)) => (case ((string_drop s54580 s54590)) of s54600 => (case ((spc_matches_prefix0 s54600)) of Some ((_, s54610)) => (case ((string_drop s54600 s54610)) of s54620 => (case ((reg_name_matches_prefix s54620 :: (( 5 Word.word * ii)) option)) of Some ((rd, s54630)) => (case ((string_drop s54620 s54630)) of s54640 => (case ((sep_matches_prefix s54640)) of Some ((_, s54650)) => (case ((string_drop s54640 s54650)) of s54660 => (case ((freg_name_matches_prefix s54660 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s54670)) => (case ((string_drop s54660 s54670)) of s1 => Some (FMV_X_W, rd, rs1, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s54580 :: " string " \ \\val _s5439_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s5439 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s5439 s54410 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s54410)) of Some ((FLE_S, s54420)) => (case ((string_drop s54410 s54420)) of s54430 => (case ((spc_matches_prefix0 s54430)) of Some ((_, s54440)) => (case ((string_drop s54430 s54440)) of s54450 => (case ((reg_name_matches_prefix s54450 :: (( 5 Word.word * ii)) option)) of Some ((rd, s54460)) => (case ((string_drop s54450 s54460)) of s54470 => (case ((sep_matches_prefix s54470)) of Some ((_, s54480)) => (case ((string_drop s54470 s54480)) of s54490 => (case ((freg_name_matches_prefix s54490 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s54500)) => (case ((string_drop s54490 s54500)) of s54510 => (case ((sep_matches_prefix s54510)) of Some ((_, s54520)) => (case ((string_drop s54510 s54520)) of s54530 => (case ((freg_name_matches_prefix s54530 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s54540)) => (case ((string_drop s54530 s54540)) of s1 => Some (FLE_S, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s54410 :: " string " \ \\val _s5422_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s5422 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s5422 s54240 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s54240)) of Some ((FLT_S, s54250)) => (case ((string_drop s54240 s54250)) of s54260 => (case ((spc_matches_prefix0 s54260)) of Some ((_, s54270)) => (case ((string_drop s54260 s54270)) of s54280 => (case ((reg_name_matches_prefix s54280 :: (( 5 Word.word * ii)) option)) of Some ((rd, s54290)) => (case ((string_drop s54280 s54290)) of s54300 => (case ((sep_matches_prefix s54300)) of Some ((_, s54310)) => (case ((string_drop s54300 s54310)) of s54320 => (case ((freg_name_matches_prefix s54320 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s54330)) => (case ((string_drop s54320 s54330)) of s54340 => (case ((sep_matches_prefix s54340)) of Some ((_, s54350)) => (case ((string_drop s54340 s54350)) of s54360 => (case ((freg_name_matches_prefix s54360 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s54370)) => (case ((string_drop s54360 s54370)) of s1 => Some (FLT_S, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s54240 :: " string " \ \\val _s5405_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s5405 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s5405 s54070 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s54070)) of Some ((FEQ_S, s54080)) => (case ((string_drop s54070 s54080)) of s54090 => (case ((spc_matches_prefix0 s54090)) of Some ((_, s54100)) => (case ((string_drop s54090 s54100)) of s54110 => (case ((reg_name_matches_prefix s54110 :: (( 5 Word.word * ii)) option)) of Some ((rd, s54120)) => (case ((string_drop s54110 s54120)) of s54130 => (case ((sep_matches_prefix s54130)) of Some ((_, s54140)) => (case ((string_drop s54130 s54140)) of s54150 => (case ((freg_name_matches_prefix s54150 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s54160)) => (case ((string_drop s54150 s54160)) of s54170 => (case ((sep_matches_prefix s54170)) of Some ((_, s54180)) => (case ((string_drop s54170 s54180)) of s54190 => (case ((freg_name_matches_prefix s54190 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s54200)) => (case ((string_drop s54190 s54200)) of s1 => Some (FEQ_S, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s54070 :: " string " \ \\val _s5388_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s5388 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s5388 s53900 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s53900)) of Some ((FMAX_S, s53910)) => (case ((string_drop s53900 s53910)) of s53920 => (case ((spc_matches_prefix0 s53920)) of Some ((_, s53930)) => (case ((string_drop s53920 s53930)) of s53940 => (case ((freg_name_matches_prefix s53940 :: (( 5 Word.word * ii)) option)) of Some ((rd, s53950)) => (case ((string_drop s53940 s53950)) of s53960 => (case ((sep_matches_prefix s53960)) of Some ((_, s53970)) => (case ((string_drop s53960 s53970)) of s53980 => (case ((freg_name_matches_prefix s53980 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s53990)) => (case ((string_drop s53980 s53990)) of s54000 => (case ((sep_matches_prefix s54000)) of Some ((_, s54010)) => (case ((string_drop s54000 s54010)) of s54020 => (case ((freg_name_matches_prefix s54020 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s54030)) => (case ((string_drop s54020 s54030)) of s1 => Some (FMAX_S, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s53900 :: " string " \ \\val _s5371_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s5371 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s5371 s53730 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s53730)) of Some ((FMIN_S, s53740)) => (case ((string_drop s53730 s53740)) of s53750 => (case ((spc_matches_prefix0 s53750)) of Some ((_, s53760)) => (case ((string_drop s53750 s53760)) of s53770 => (case ((freg_name_matches_prefix s53770 :: (( 5 Word.word * ii)) option)) of Some ((rd, s53780)) => (case ((string_drop s53770 s53780)) of s53790 => (case ((sep_matches_prefix s53790)) of Some ((_, s53800)) => (case ((string_drop s53790 s53800)) of s53810 => (case ((freg_name_matches_prefix s53810 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s53820)) => (case ((string_drop s53810 s53820)) of s53830 => (case ((sep_matches_prefix s53830)) of Some ((_, s53840)) => (case ((string_drop s53830 s53840)) of s53850 => (case ((freg_name_matches_prefix s53850 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s53860)) => (case ((string_drop s53850 s53860)) of s1 => Some (FMIN_S, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s53730 :: " string " \ \\val _s5354_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s5354 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s5354 s53560 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s53560)) of Some ((FSGNJX_S, s53570)) => (case ((string_drop s53560 s53570)) of s53580 => (case ((spc_matches_prefix0 s53580)) of Some ((_, s53590)) => (case ((string_drop s53580 s53590)) of s53600 => (case ((freg_name_matches_prefix s53600 :: (( 5 Word.word * ii)) option)) of Some ((rd, s53610)) => (case ((string_drop s53600 s53610)) of s53620 => (case ((sep_matches_prefix s53620)) of Some ((_, s53630)) => (case ((string_drop s53620 s53630)) of s53640 => (case ((freg_name_matches_prefix s53640 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s53650)) => (case ((string_drop s53640 s53650)) of s53660 => (case ((sep_matches_prefix s53660)) of Some ((_, s53670)) => (case ((string_drop s53660 s53670)) of s53680 => (case ((freg_name_matches_prefix s53680 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s53690)) => (case ((string_drop s53680 s53690)) of s1 => Some (FSGNJX_S, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s53560 :: " string " \ \\val _s5337_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s5337 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s5337 s53390 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s53390)) of Some ((FSGNJN_S, s53400)) => (case ((string_drop s53390 s53400)) of s53410 => (case ((spc_matches_prefix0 s53410)) of Some ((_, s53420)) => (case ((string_drop s53410 s53420)) of s53430 => (case ((freg_name_matches_prefix s53430 :: (( 5 Word.word * ii)) option)) of Some ((rd, s53440)) => (case ((string_drop s53430 s53440)) of s53450 => (case ((sep_matches_prefix s53450)) of Some ((_, s53460)) => (case ((string_drop s53450 s53460)) of s53470 => (case ((freg_name_matches_prefix s53470 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s53480)) => (case ((string_drop s53470 s53480)) of s53490 => (case ((sep_matches_prefix s53490)) of Some ((_, s53500)) => (case ((string_drop s53490 s53500)) of s53510 => (case ((freg_name_matches_prefix s53510 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s53520)) => (case ((string_drop s53510 s53520)) of s1 => Some (FSGNJN_S, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s53390 :: " string " \ \\val _s5320_ : string -> maybe ((f_bin_op_S * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s5320 :: \ string \(f_bin_op_S*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s5320 s53220 = ( (case ((f_bin_type_mnemonic_S_matches_prefix s53220)) of Some ((FSGNJ_S, s53230)) => (case ((string_drop s53220 s53230)) of s53240 => (case ((spc_matches_prefix0 s53240)) of Some ((_, s53250)) => (case ((string_drop s53240 s53250)) of s53260 => (case ((freg_name_matches_prefix s53260 :: (( 5 Word.word * ii)) option)) of Some ((rd, s53270)) => (case ((string_drop s53260 s53270)) of s53280 => (case ((sep_matches_prefix s53280)) of Some ((_, s53290)) => (case ((string_drop s53280 s53290)) of s53300 => (case ((freg_name_matches_prefix s53300 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s53310)) => (case ((string_drop s53300 s53310)) of s53320 => (case ((sep_matches_prefix s53320)) of Some ((_, s53330)) => (case ((string_drop s53320 s53330)) of s53340 => (case ((freg_name_matches_prefix s53340 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s53350)) => (case ((string_drop s53340 s53350)) of s1 => Some (FSGNJ_S, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s53220 :: " string " \ \\val _s5303_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5303 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5303 s53050 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s53050)) of Some ((FCVT_S_LU, s53060)) => (case ((string_drop s53050 s53060)) of s53070 => (case ((spc_matches_prefix0 s53070)) of Some ((_, s53080)) => (case ((string_drop s53070 s53080)) of s53090 => (case ((freg_name_matches_prefix s53090 :: (( 5 Word.word * ii)) option)) of Some ((rd, s53100)) => (case ((string_drop s53090 s53100)) of s53110 => (case ((sep_matches_prefix s53110)) of Some ((_, s53120)) => (case ((string_drop s53110 s53120)) of s53130 => (case ((reg_name_matches_prefix s53130 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s53140)) => (case ((string_drop s53130 s53140)) of s53150 => (case ((sep_matches_prefix s53150)) of Some ((_, s53160)) => (case ((string_drop s53150 s53160)) of s53170 => (case ((frm_mnemonic_matches_prefix s53170)) of Some ((rm, s53180)) => (case ((string_drop s53170 s53180)) of s1 => Some (FCVT_S_LU, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s53050 :: " string " \ \\val _s5286_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5286 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5286 s52880 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s52880)) of Some ((FCVT_S_L, s52890)) => (case ((string_drop s52880 s52890)) of s52900 => (case ((spc_matches_prefix0 s52900)) of Some ((_, s52910)) => (case ((string_drop s52900 s52910)) of s52920 => (case ((freg_name_matches_prefix s52920 :: (( 5 Word.word * ii)) option)) of Some ((rd, s52930)) => (case ((string_drop s52920 s52930)) of s52940 => (case ((sep_matches_prefix s52940)) of Some ((_, s52950)) => (case ((string_drop s52940 s52950)) of s52960 => (case ((reg_name_matches_prefix s52960 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s52970)) => (case ((string_drop s52960 s52970)) of s52980 => (case ((sep_matches_prefix s52980)) of Some ((_, s52990)) => (case ((string_drop s52980 s52990)) of s53000 => (case ((frm_mnemonic_matches_prefix s53000)) of Some ((rm, s53010)) => (case ((string_drop s53000 s53010)) of s1 => Some (FCVT_S_L, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s52880 :: " string " \ \\val _s5269_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5269 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5269 s52710 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s52710)) of Some ((FCVT_LU_S, s52720)) => (case ((string_drop s52710 s52720)) of s52730 => (case ((spc_matches_prefix0 s52730)) of Some ((_, s52740)) => (case ((string_drop s52730 s52740)) of s52750 => (case ((reg_name_matches_prefix s52750 :: (( 5 Word.word * ii)) option)) of Some ((rd, s52760)) => (case ((string_drop s52750 s52760)) of s52770 => (case ((sep_matches_prefix s52770)) of Some ((_, s52780)) => (case ((string_drop s52770 s52780)) of s52790 => (case ((freg_name_matches_prefix s52790 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s52800)) => (case ((string_drop s52790 s52800)) of s52810 => (case ((sep_matches_prefix s52810)) of Some ((_, s52820)) => (case ((string_drop s52810 s52820)) of s52830 => (case ((frm_mnemonic_matches_prefix s52830)) of Some ((rm, s52840)) => (case ((string_drop s52830 s52840)) of s1 => Some (FCVT_LU_S, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s52710 :: " string " \ \\val _s5252_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5252 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5252 s52540 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s52540)) of Some ((FCVT_L_S, s52550)) => (case ((string_drop s52540 s52550)) of s52560 => (case ((spc_matches_prefix0 s52560)) of Some ((_, s52570)) => (case ((string_drop s52560 s52570)) of s52580 => (case ((reg_name_matches_prefix s52580 :: (( 5 Word.word * ii)) option)) of Some ((rd, s52590)) => (case ((string_drop s52580 s52590)) of s52600 => (case ((sep_matches_prefix s52600)) of Some ((_, s52610)) => (case ((string_drop s52600 s52610)) of s52620 => (case ((freg_name_matches_prefix s52620 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s52630)) => (case ((string_drop s52620 s52630)) of s52640 => (case ((sep_matches_prefix s52640)) of Some ((_, s52650)) => (case ((string_drop s52640 s52650)) of s52660 => (case ((frm_mnemonic_matches_prefix s52660)) of Some ((rm, s52670)) => (case ((string_drop s52660 s52670)) of s1 => Some (FCVT_L_S, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s52540 :: " string " \ \\val _s5235_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5235 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5235 s52370 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s52370)) of Some ((FCVT_S_WU, s52380)) => (case ((string_drop s52370 s52380)) of s52390 => (case ((spc_matches_prefix0 s52390)) of Some ((_, s52400)) => (case ((string_drop s52390 s52400)) of s52410 => (case ((freg_name_matches_prefix s52410 :: (( 5 Word.word * ii)) option)) of Some ((rd, s52420)) => (case ((string_drop s52410 s52420)) of s52430 => (case ((sep_matches_prefix s52430)) of Some ((_, s52440)) => (case ((string_drop s52430 s52440)) of s52450 => (case ((reg_name_matches_prefix s52450 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s52460)) => (case ((string_drop s52450 s52460)) of s52470 => (case ((sep_matches_prefix s52470)) of Some ((_, s52480)) => (case ((string_drop s52470 s52480)) of s52490 => (case ((frm_mnemonic_matches_prefix s52490)) of Some ((rm, s52500)) => (case ((string_drop s52490 s52500)) of s1 => Some (FCVT_S_WU, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s52370 :: " string " \ \\val _s5218_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5218 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5218 s52200 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s52200)) of Some ((FCVT_S_W, s52210)) => (case ((string_drop s52200 s52210)) of s52220 => (case ((spc_matches_prefix0 s52220)) of Some ((_, s52230)) => (case ((string_drop s52220 s52230)) of s52240 => (case ((freg_name_matches_prefix s52240 :: (( 5 Word.word * ii)) option)) of Some ((rd, s52250)) => (case ((string_drop s52240 s52250)) of s52260 => (case ((sep_matches_prefix s52260)) of Some ((_, s52270)) => (case ((string_drop s52260 s52270)) of s52280 => (case ((reg_name_matches_prefix s52280 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s52290)) => (case ((string_drop s52280 s52290)) of s52300 => (case ((sep_matches_prefix s52300)) of Some ((_, s52310)) => (case ((string_drop s52300 s52310)) of s52320 => (case ((frm_mnemonic_matches_prefix s52320)) of Some ((rm, s52330)) => (case ((string_drop s52320 s52330)) of s1 => Some (FCVT_S_W, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s52200 :: " string " \ \\val _s5201_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5201 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5201 s52030 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s52030)) of Some ((FCVT_WU_S, s52040)) => (case ((string_drop s52030 s52040)) of s52050 => (case ((spc_matches_prefix0 s52050)) of Some ((_, s52060)) => (case ((string_drop s52050 s52060)) of s52070 => (case ((reg_name_matches_prefix s52070 :: (( 5 Word.word * ii)) option)) of Some ((rd, s52080)) => (case ((string_drop s52070 s52080)) of s52090 => (case ((sep_matches_prefix s52090)) of Some ((_, s52100)) => (case ((string_drop s52090 s52100)) of s52110 => (case ((freg_name_matches_prefix s52110 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s52120)) => (case ((string_drop s52110 s52120)) of s52130 => (case ((sep_matches_prefix s52130)) of Some ((_, s52140)) => (case ((string_drop s52130 s52140)) of s52150 => (case ((frm_mnemonic_matches_prefix s52150)) of Some ((rm, s52160)) => (case ((string_drop s52150 s52160)) of s1 => Some (FCVT_WU_S, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s52030 :: " string " \ \\val _s5184_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5184 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5184 s51860 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s51860)) of Some ((FCVT_W_S, s51870)) => (case ((string_drop s51860 s51870)) of s51880 => (case ((spc_matches_prefix0 s51880)) of Some ((_, s51890)) => (case ((string_drop s51880 s51890)) of s51900 => (case ((reg_name_matches_prefix s51900 :: (( 5 Word.word * ii)) option)) of Some ((rd, s51910)) => (case ((string_drop s51900 s51910)) of s51920 => (case ((sep_matches_prefix s51920)) of Some ((_, s51930)) => (case ((string_drop s51920 s51930)) of s51940 => (case ((freg_name_matches_prefix s51940 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s51950)) => (case ((string_drop s51940 s51950)) of s51960 => (case ((sep_matches_prefix s51960)) of Some ((_, s51970)) => (case ((string_drop s51960 s51970)) of s51980 => (case ((frm_mnemonic_matches_prefix s51980)) of Some ((rm, s51990)) => (case ((string_drop s51980 s51990)) of s1 => Some (FCVT_W_S, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s51860 :: " string " \ \\val _s5167_ : string -> maybe ((f_un_rm_op_S * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5167 :: \ string \(f_un_rm_op_S*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5167 s51690 = ( (case ((f_un_rm_type_mnemonic_S_matches_prefix s51690)) of Some ((FSQRT_S, s51700)) => (case ((string_drop s51690 s51700)) of s51710 => (case ((spc_matches_prefix0 s51710)) of Some ((_, s51720)) => (case ((string_drop s51710 s51720)) of s51730 => (case ((freg_name_matches_prefix s51730 :: (( 5 Word.word * ii)) option)) of Some ((rd, s51740)) => (case ((string_drop s51730 s51740)) of s51750 => (case ((sep_matches_prefix s51750)) of Some ((_, s51760)) => (case ((string_drop s51750 s51760)) of s51770 => (case ((freg_name_matches_prefix s51770 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s51780)) => (case ((string_drop s51770 s51780)) of s51790 => (case ((sep_matches_prefix s51790)) of Some ((_, s51800)) => (case ((string_drop s51790 s51800)) of s51810 => (case ((frm_mnemonic_matches_prefix s51810)) of Some ((rm, s51820)) => (case ((string_drop s51810 s51820)) of s1 => Some (FSQRT_S, rd, rs1, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s51690 :: " string " \ \\val _s5146_ : string -> maybe ((f_bin_rm_op_S * mword ty5 * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5146 :: \ string \(f_bin_rm_op_S*(5)Word.word*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5146 s51480 = ( (case ((f_bin_rm_type_mnemonic_S_matches_prefix s51480)) of Some ((op1, s51490)) => (case ((string_drop s51480 s51490)) of s51500 => (case ((spc_matches_prefix0 s51500)) of Some ((_, s51510)) => (case ((string_drop s51500 s51510)) of s51520 => (case ((freg_name_matches_prefix s51520 :: (( 5 Word.word * ii)) option)) of Some ((rd, s51530)) => (case ((string_drop s51520 s51530)) of s51540 => (case ((sep_matches_prefix s51540)) of Some ((_, s51550)) => (case ((string_drop s51540 s51550)) of s51560 => (case ((freg_name_matches_prefix s51560 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s51570)) => (case ((string_drop s51560 s51570)) of s51580 => (case ((sep_matches_prefix s51580)) of Some ((_, s51590)) => (case ((string_drop s51580 s51590)) of s51600 => (case ((freg_name_matches_prefix s51600 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s51610)) => (case ((string_drop s51600 s51610)) of s51620 => (case ((sep_matches_prefix s51620)) of Some ((_, s51630)) => (case ((string_drop s51620 s51630)) of s51640 => (case ((frm_mnemonic_matches_prefix s51640)) of Some ((rm, s51650)) => (case ((string_drop s51640 s51650)) of s1 => Some (op1, rd, rs1, rs2, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s51480 :: " string " \ \\val _s5121_ : string -> maybe ((f_madd_op_S * mword ty5 * mword ty5 * mword ty5 * mword ty5 * rounding_mode * string))\\ definition s5121 :: \ string \(f_madd_op_S*(5)Word.word*(5)Word.word*(5)Word.word*(5)Word.word*rounding_mode*string)option \ where \ s5121 s51230 = ( (case ((f_madd_type_mnemonic_S_matches_prefix s51230)) of Some ((op1, s51240)) => (case ((string_drop s51230 s51240)) of s51250 => (case ((spc_matches_prefix0 s51250)) of Some ((_, s51260)) => (case ((string_drop s51250 s51260)) of s51270 => (case ((freg_name_matches_prefix s51270 :: (( 5 Word.word * ii)) option)) of Some ((rd, s51280)) => (case ((string_drop s51270 s51280)) of s51290 => (case ((sep_matches_prefix s51290)) of Some ((_, s51300)) => (case ((string_drop s51290 s51300)) of s51310 => (case ((freg_name_matches_prefix s51310 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s51320)) => (case ((string_drop s51310 s51320)) of s51330 => (case ((sep_matches_prefix s51330)) of Some ((_, s51340)) => (case ((string_drop s51330 s51340)) of s51350 => (case ((freg_name_matches_prefix s51350 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s51360)) => (case ((string_drop s51350 s51360)) of s51370 => (case ((sep_matches_prefix s51370)) of Some ((_, s51380)) => (case ((string_drop s51370 s51380)) of s51390 => (case ((freg_name_matches_prefix s51390 :: (( 5 Word.word * ii)) option)) of Some ((rs3, s51400)) => (case ((string_drop s51390 s51400)) of s51410 => (case ((sep_matches_prefix s51410)) of Some ((_, s51420)) => (case ((string_drop s51410 s51420)) of s51430 => (case ((frm_mnemonic_matches_prefix s51430)) of Some ((rm, s51440)) => (case ((string_drop s51430 s51440)) of s1 => Some (op1, rd, rs1, rs2, rs3, rm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s51230 :: " string " \ \\val _s5097_ : string -> maybe ((word_width * mword ty5 * mword ty12 * mword ty5 * string))\\ definition s5097 :: \ string \(word_width*(5)Word.word*(12)Word.word*(5)Word.word*string)option \ where \ s5097 s50980 = ( (let s50990 = s50980 in if ((string_startswith s50990 (''fs''))) then (case ((string_drop s50990 ((string_length (''fs''))))) of s51000 => (case ((size_mnemonic_matches_prefix s51000)) of Some ((width, s51010)) => (case ((string_drop s51000 s51010)) of s51020 => (case ((spc_matches_prefix0 s51020)) of Some ((_, s51030)) => (case ((string_drop s51020 s51030)) of s51040 => (case ((freg_name_matches_prefix s51040 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s51050)) => (case ((string_drop s51040 s51050)) of s51060 => (case ((sep_matches_prefix s51060)) of Some ((_, s51070)) => (case ((string_drop s51060 s51070)) of s51080 => (case ((hex_bits_12_matches_prefix0 s51080 :: (( 12 Word.word * ii)) option)) of Some ((imm, s51090)) => (case ((string_drop s51080 s51090)) of s51100 => (case ((opt_spc_matches_prefix0 s51100)) of Some ((_, s51110)) => (let s51120 = (string_drop s51100 s51110) in if ((string_startswith s51120 (''(''))) then (case ((string_drop s51120 ((string_length (''(''))))) of s51130 => (case ((opt_spc_matches_prefix0 s51130)) of Some ((_, s51140)) => (case ((string_drop s51130 s51140)) of s51150 => (case ((reg_name_matches_prefix s51150 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s51160)) => (case ((string_drop s51150 s51160)) of s51170 => (case ((opt_spc_matches_prefix0 s51170)) of Some ((_, s51180)) => (let s51190 = (string_drop s51170 s51180) in if ((string_startswith s51190 ('')''))) then (case ((string_drop s51190 ((string_length ('')''))))) of s1 => Some (width, rs2, imm, rs1, s1) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s50980 :: " string " \ \\val _s5073_ : string -> maybe ((word_width * mword ty5 * mword ty12 * mword ty5 * string))\\ definition s5073 :: \ string \(word_width*(5)Word.word*(12)Word.word*(5)Word.word*string)option \ where \ s5073 s50740 = ( (let s50750 = s50740 in if ((string_startswith s50750 (''fl''))) then (case ((string_drop s50750 ((string_length (''fl''))))) of s50760 => (case ((size_mnemonic_matches_prefix s50760)) of Some ((width, s50770)) => (case ((string_drop s50760 s50770)) of s50780 => (case ((spc_matches_prefix0 s50780)) of Some ((_, s50790)) => (case ((string_drop s50780 s50790)) of s50800 => (case ((freg_name_matches_prefix s50800 :: (( 5 Word.word * ii)) option)) of Some ((rd, s50810)) => (case ((string_drop s50800 s50810)) of s50820 => (case ((sep_matches_prefix s50820)) of Some ((_, s50830)) => (case ((string_drop s50820 s50830)) of s50840 => (case ((hex_bits_12_matches_prefix0 s50840 :: (( 12 Word.word * ii)) option)) of Some ((imm, s50850)) => (case ((string_drop s50840 s50850)) of s50860 => (case ((opt_spc_matches_prefix0 s50860)) of Some ((_, s50870)) => (let s50880 = (string_drop s50860 s50870) in if ((string_startswith s50880 (''(''))) then (case ((string_drop s50880 ((string_length (''(''))))) of s50890 => (case ((opt_spc_matches_prefix0 s50890)) of Some ((_, s50900)) => (case ((string_drop s50890 s50900)) of s50910 => (case ((reg_name_matches_prefix s50910 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s50920)) => (case ((string_drop s50910 s50920)) of s50930 => (case ((opt_spc_matches_prefix0 s50930)) of Some ((_, s50940)) => (let s50950 = (string_drop s50930 s50940) in if ((string_startswith s50950 ('')''))) then (case ((string_drop s50950 ((string_length ('')''))))) of s1 => Some (width, rd, imm, rs1, s1) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s50740 :: " string " \ \\val _s5061_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12 * string))\\ definition s5061 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word*string)option \ where \ s5061 s50620 = ( (let s50630 = s50620 in if ((string_startswith s50630 (''fence.i.reserved.''))) then (case ((string_drop s50630 ((string_length (''fence.i.reserved.''))))) of s50640 => (case ((reg_name_matches_prefix s50640 :: (( 5 Word.word * ii)) option)) of Some ((rd, s50650)) => (let s50660 = (string_drop s50640 s50650) in if ((string_startswith s50660 (''.''))) then (case ((string_drop s50660 ((string_length (''.''))))) of s50670 => (case ((reg_name_matches_prefix s50670 :: (( 5 Word.word * ii)) option)) of Some ((rs, s50680)) => (let s50690 = (string_drop s50670 s50680) in if ((string_startswith s50690 (''.''))) then (case ((string_drop s50690 ((string_length (''.''))))) of s50700 => (case ((hex_bits_12_matches_prefix0 s50700 :: (( 12 Word.word * ii)) option)) of Some ((imm, s50710)) => (case ((string_drop s50700 s50710)) of s1 => Some (rd, rs, imm, s1) ) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None))\ for s50620 :: " string " \ \\val _s5043_ : string -> maybe ((mword ty4 * mword ty4 * mword ty5 * mword ty5 * mword ty4 * string))\\ definition s5043 :: \ string \((4)Word.word*(4)Word.word*(5)Word.word*(5)Word.word*(4)Word.word*string)option \ where \ s5043 s50440 = ( (let s50450 = s50440 in if ((string_startswith s50450 (''fence.reserved.''))) then (case ((string_drop s50450 ((string_length (''fence.reserved.''))))) of s50460 => (case ((fence_bits_matches_prefix s50460 :: (( 4 Word.word * ii)) option)) of Some ((pred, s50470)) => (let s50480 = (string_drop s50460 s50470) in if ((string_startswith s50480 (''.''))) then (case ((string_drop s50480 ((string_length (''.''))))) of s50490 => (case ((fence_bits_matches_prefix s50490 :: (( 4 Word.word * ii)) option)) of Some ((succ, s50500)) => (let s50510 = (string_drop s50490 s50500) in if ((string_startswith s50510 (''.''))) then (case ((string_drop s50510 ((string_length (''.''))))) of s50520 => (case ((reg_name_matches_prefix s50520 :: (( 5 Word.word * ii)) option)) of Some ((rs, s50530)) => (let s50540 = (string_drop s50520 s50530) in if ((string_startswith s50540 (''.''))) then (case ((string_drop s50540 ((string_length (''.''))))) of s50550 => (case ((reg_name_matches_prefix s50550 :: (( 5 Word.word * ii)) option)) of Some ((rd, s50560)) => (let s50570 = (string_drop s50550 s50560) in if ((string_startswith s50570 (''.''))) then (case ((string_drop s50570 ((string_length (''.''))))) of s50580 => (case ((hex_bits_4_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s50580 :: (( 4 Word.word * ii)) option)) of Some ((fm, s50590)) => (case ((string_drop s50580 s50590)) of s1 => Some (pred, succ, rs, rd, fm, s1) ) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None) | _ => None ) ) else None))\ for s50440 :: " string " \ \\val _s5037_ : string -> maybe ((mword ty3 * string))\\ definition s5037 :: \ string \((3)Word.word*string)option \ where \ s5037 s50380 = ( (let s50390 = s50380 in if ((string_startswith s50390 (''c.srai.hint.''))) then (case ((string_drop s50390 ((string_length (''c.srai.hint.''))))) of s50400 => (case ((creg_name_matches_prefix s50400 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s50410)) => (case ((string_drop s50400 s50410)) of s1 => Some (rsd, s1) ) | _ => None ) ) else None))\ for s50380 :: " string " \ \\val _s5031_ : string -> maybe ((mword ty3 * string))\\ definition s5031 :: \ string \((3)Word.word*string)option \ where \ s5031 s50320 = ( (let s50330 = s50320 in if ((string_startswith s50330 (''c.srli.hint.''))) then (case ((string_drop s50330 ((string_length (''c.srli.hint.''))))) of s50340 => (case ((creg_name_matches_prefix s50340 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s50350)) => (case ((string_drop s50340 s50350)) of s1 => Some (rsd, s1) ) | _ => None ) ) else None))\ for s50320 :: " string " \ \\val _s5022_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s5022 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s5022 s50230 = ( (let s50240 = s50230 in if ((string_startswith s50240 (''c.slli.hint.''))) then (case ((string_drop s50240 ((string_length (''c.slli.hint.''))))) of s50250 => (case ((reg_name_matches_prefix s50250 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s50260)) => (let s50270 = (string_drop s50250 s50260) in if ((string_startswith s50270 (''.''))) then (case ((string_drop s50270 ((string_length (''.''))))) of s50280 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s50280 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s50290)) => (case ((string_drop s50280 s50290)) of s1 => Some (rsd, shamt, s1) ) | _ => None ) ) else None) | _ => None ) ) else None))\ for s50230 :: " string " \ \\val _s5016_ : string -> maybe ((mword ty5 * string))\\ definition s5016 :: \ string \((5)Word.word*string)option \ where \ s5016 s50170 = ( (let s50180 = s50170 in if ((string_startswith s50180 (''c.add.hint.''))) then (case ((string_drop s50180 ((string_length (''c.add.hint.''))))) of s50190 => (case ((reg_name_matches_prefix s50190 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s50200)) => (case ((string_drop s50190 s50200)) of s1 => Some (rs2, s1) ) | _ => None ) ) else None))\ for s50170 :: " string " \ \\val _s5010_ : string -> maybe ((mword ty5 * string))\\ definition s5010 :: \ string \((5)Word.word*string)option \ where \ s5010 s50110 = ( (let s50120 = s50110 in if ((string_startswith s50120 (''c.mv.hint.''))) then (case ((string_drop s50120 ((string_length (''c.mv.hint.''))))) of s50130 => (case ((reg_name_matches_prefix s50130 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s50140)) => (case ((string_drop s50130 s50140)) of s1 => Some (rs2, s1) ) | _ => None ) ) else None))\ for s50110 :: " string " \ \\val _s5004_ : string -> maybe ((mword ty6 * string))\\ definition s5004 :: \ string \((6)Word.word*string)option \ where \ s5004 s50050 = ( (let s50060 = s50050 in if ((string_startswith s50060 (''c.lui.hint.''))) then (case ((string_drop s50060 ((string_length (''c.lui.hint.''))))) of s50070 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s50070 :: (( 6 Word.word * ii)) option)) of Some ((imm, s50080)) => (case ((string_drop s50070 s50080)) of s1 => Some (imm, s1) ) | _ => None ) ) else None))\ for s50050 :: " string " \ \\val _s4998_ : string -> maybe ((mword ty6 * string))\\ definition s4998 :: \ string \((6)Word.word*string)option \ where \ s4998 s49990 = ( (let s50000 = s49990 in if ((string_startswith s50000 (''c.li.hint.''))) then (case ((string_drop s50000 ((string_length (''c.li.hint.''))))) of s50010 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s50010 :: (( 6 Word.word * ii)) option)) of Some ((imm, s50020)) => (case ((string_drop s50010 s50020)) of s1 => Some (imm, s1) ) | _ => None ) ) else None))\ for s49990 :: " string " \ \\val _s4992_ : string -> maybe ((mword ty5 * string))\\ definition s4992 :: \ string \((5)Word.word*string)option \ where \ s4992 s49930 = ( (let s49940 = s49930 in if ((string_startswith s49940 (''c.addi.hint.''))) then (case ((string_drop s49940 ((string_length (''c.addi.hint.''))))) of s49950 => (case ((reg_name_matches_prefix s49950 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s49960)) => (case ((string_drop s49950 s49960)) of s1 => Some (rsd, s1) ) | _ => None ) ) else None))\ for s49930 :: " string " \ \\val _s4986_ : string -> maybe ((mword ty6 * string))\\ definition s4986 :: \ string \((6)Word.word*string)option \ where \ s4986 s49870 = ( (let s49880 = s49870 in if ((string_startswith s49880 (''c.nop.hint.''))) then (case ((string_drop s49880 ((string_length (''c.nop.hint.''))))) of s49890 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s49890 :: (( 6 Word.word * ii)) option)) of Some ((imm, s49900)) => (case ((string_drop s49890 s49900)) of s1 => Some (imm, s1) ) | _ => None ) ) else None))\ for s49870 :: " string " \ \\val _s4982_ : string -> maybe string\\ definition s4982 :: \ string \(string)option \ where \ s4982 s49830 = ( (let s49840 = s49830 in if ((string_startswith s49840 (''uret''))) then (case ((string_drop s49840 ((string_length (''uret''))))) of s1 => Some s1 ) else None))\ for s49830 :: " string " \ \\val _s4965_ : string -> maybe ((csrop * mword ty5 * mword ty12 * mword ty5 * string))\\ definition s4965 :: \ string \(csrop*(5)Word.word*(12)Word.word*(5)Word.word*string)option \ where \ s4965 s49670 = ( (case ((csr_mnemonic_matches_prefix s49670)) of Some ((op1, s49680)) => (case ((string_drop s49670 s49680)) of s49690 => (case ((spc_matches_prefix0 s49690)) of Some ((_, s49700)) => (case ((string_drop s49690 s49700)) of s49710 => (case ((reg_name_matches_prefix s49710 :: (( 5 Word.word * ii)) option)) of Some ((rd, s49720)) => (case ((string_drop s49710 s49720)) of s49730 => (case ((sep_matches_prefix s49730)) of Some ((_, s49740)) => (case ((string_drop s49730 s49740)) of s49750 => (case ((csr_name_map_matches_prefix s49750 :: (( 12 Word.word * ii)) option)) of Some ((csr, s49760)) => (case ((string_drop s49750 s49760)) of s49770 => (case ((sep_matches_prefix s49770)) of Some ((_, s49780)) => (case ((string_drop s49770 s49780)) of s49790 => (case ((reg_name_matches_prefix s49790 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s49800)) => (case ((string_drop s49790 s49800)) of s1 => Some (op1, rd, csr, rs1, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s49670 :: " string " \ \\val _s4947_ : string -> maybe ((csrop * mword ty5 * mword ty12 * mword ty5 * string))\\ definition s4947 :: \ string \(csrop*(5)Word.word*(12)Word.word*(5)Word.word*string)option \ where \ s4947 s49490 = ( (case ((csr_mnemonic_matches_prefix s49490)) of Some ((op1, s49500)) => (let s49510 = (string_drop s49490 s49500) in if ((string_startswith s49510 (''i''))) then (case ((string_drop s49510 ((string_length (''i''))))) of s49520 => (case ((spc_matches_prefix0 s49520)) of Some ((_, s49530)) => (case ((string_drop s49520 s49530)) of s49540 => (case ((reg_name_matches_prefix s49540 :: (( 5 Word.word * ii)) option)) of Some ((rd, s49550)) => (case ((string_drop s49540 s49550)) of s49560 => (case ((sep_matches_prefix s49560)) of Some ((_, s49570)) => (case ((string_drop s49560 s49570)) of s49580 => (case ((csr_name_map_matches_prefix s49580 :: (( 12 Word.word * ii)) option)) of Some ((csr, s49590)) => (case ((string_drop s49580 s49590)) of s49600 => (case ((sep_matches_prefix s49600)) of Some ((_, s49610)) => (case ((string_drop s49600 s49610)) of s49620 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s49620 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s49630)) => (case ((string_drop s49620 s49630)) of s1 => Some (op1, rd, csr, rs1, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ))\ for s49490 :: " string " \ \\val _s4928_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4928 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4928 s49290 = ( (let s49300 = s49290 in if ((string_startswith s49300 (''rem''))) then (case ((string_drop s49300 ((string_length (''rem''))))) of s49310 => (case ((maybe_not_u_matches_prefix s49310)) of Some ((s, s49320)) => (let s49330 = (string_drop s49310 s49320) in if ((string_startswith s49330 (''w''))) then (case ((string_drop s49330 ((string_length (''w''))))) of s49340 => (case ((spc_matches_prefix0 s49340)) of Some ((_, s49350)) => (case ((string_drop s49340 s49350)) of s49360 => (case ((reg_name_matches_prefix s49360 :: (( 5 Word.word * ii)) option)) of Some ((rd, s49370)) => (case ((string_drop s49360 s49370)) of s49380 => (case ((sep_matches_prefix s49380)) of Some ((_, s49390)) => (case ((string_drop s49380 s49390)) of s49400 => (case ((reg_name_matches_prefix s49400 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s49410)) => (case ((string_drop s49400 s49410)) of s49420 => (case ((sep_matches_prefix s49420)) of Some ((_, s49430)) => (case ((string_drop s49420 s49430)) of s49440 => (case ((reg_name_matches_prefix s49440 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s49450)) => (case ((string_drop s49440 s49450)) of s2 => Some (s, rd, rs1, rs2, s2) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) else None))\ for s49290 :: " string " \ \\val _s4909_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4909 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4909 s49100 = ( (let s49110 = s49100 in if ((string_startswith s49110 (''div''))) then (case ((string_drop s49110 ((string_length (''div''))))) of s49120 => (case ((maybe_not_u_matches_prefix s49120)) of Some ((s, s49130)) => (let s49140 = (string_drop s49120 s49130) in if ((string_startswith s49140 (''w''))) then (case ((string_drop s49140 ((string_length (''w''))))) of s49150 => (case ((spc_matches_prefix0 s49150)) of Some ((_, s49160)) => (case ((string_drop s49150 s49160)) of s49170 => (case ((reg_name_matches_prefix s49170 :: (( 5 Word.word * ii)) option)) of Some ((rd, s49180)) => (case ((string_drop s49170 s49180)) of s49190 => (case ((sep_matches_prefix s49190)) of Some ((_, s49200)) => (case ((string_drop s49190 s49200)) of s49210 => (case ((reg_name_matches_prefix s49210 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s49220)) => (case ((string_drop s49210 s49220)) of s49230 => (case ((sep_matches_prefix s49230)) of Some ((_, s49240)) => (case ((string_drop s49230 s49240)) of s49250 => (case ((reg_name_matches_prefix s49250 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s49260)) => (case ((string_drop s49250 s49260)) of s2 => Some (s, rd, rs1, rs2, s2) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) else None))\ for s49100 :: " string " \ \\val _s4893_ : string -> maybe ((mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4893 :: \ string \((5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4893 s48940 = ( (let s48950 = s48940 in if ((string_startswith s48950 (''mulw''))) then (case ((string_drop s48950 ((string_length (''mulw''))))) of s48960 => (case ((spc_matches_prefix0 s48960)) of Some ((_, s48970)) => (case ((string_drop s48960 s48970)) of s48980 => (case ((reg_name_matches_prefix s48980 :: (( 5 Word.word * ii)) option)) of Some ((rd, s48990)) => (case ((string_drop s48980 s48990)) of s49000 => (case ((sep_matches_prefix s49000)) of Some ((_, s49010)) => (case ((string_drop s49000 s49010)) of s49020 => (case ((reg_name_matches_prefix s49020 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s49030)) => (case ((string_drop s49020 s49030)) of s49040 => (case ((sep_matches_prefix s49040)) of Some ((_, s49050)) => (case ((string_drop s49040 s49050)) of s49060 => (case ((reg_name_matches_prefix s49060 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s49070)) => (case ((string_drop s49060 s49070)) of s1 => Some (rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s48940 :: " string " \ \\val _s4875_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4875 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4875 s48760 = ( (let s48770 = s48760 in if ((string_startswith s48770 (''rem''))) then (case ((string_drop s48770 ((string_length (''rem''))))) of s48780 => (case ((maybe_not_u_matches_prefix s48780)) of Some ((s, s48790)) => (case ((string_drop s48780 s48790)) of s48800 => (case ((spc_matches_prefix0 s48800)) of Some ((_, s48810)) => (case ((string_drop s48800 s48810)) of s48820 => (case ((reg_name_matches_prefix s48820 :: (( 5 Word.word * ii)) option)) of Some ((rd, s48830)) => (case ((string_drop s48820 s48830)) of s48840 => (case ((sep_matches_prefix s48840)) of Some ((_, s48850)) => (case ((string_drop s48840 s48850)) of s48860 => (case ((reg_name_matches_prefix s48860 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s48870)) => (case ((string_drop s48860 s48870)) of s48880 => (case ((sep_matches_prefix s48880)) of Some ((_, s48890)) => (case ((string_drop s48880 s48890)) of s48900 => (case ((reg_name_matches_prefix s48900 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s48910)) => (case ((string_drop s48900 s48910)) of s2 => Some (s, rd, rs1, rs2, s2) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s48760 :: " string " \ \\val _s4857_ : string -> maybe ((bool * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4857 :: \ string \(bool*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4857 s48580 = ( (let s48590 = s48580 in if ((string_startswith s48590 (''div''))) then (case ((string_drop s48590 ((string_length (''div''))))) of s48600 => (case ((maybe_not_u_matches_prefix s48600)) of Some ((s, s48610)) => (case ((string_drop s48600 s48610)) of s48620 => (case ((spc_matches_prefix0 s48620)) of Some ((_, s48630)) => (case ((string_drop s48620 s48630)) of s48640 => (case ((reg_name_matches_prefix s48640 :: (( 5 Word.word * ii)) option)) of Some ((rd, s48650)) => (case ((string_drop s48640 s48650)) of s48660 => (case ((sep_matches_prefix s48660)) of Some ((_, s48670)) => (case ((string_drop s48660 s48670)) of s48680 => (case ((reg_name_matches_prefix s48680 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s48690)) => (case ((string_drop s48680 s48690)) of s48700 => (case ((sep_matches_prefix s48700)) of Some ((_, s48710)) => (case ((string_drop s48700 s48710)) of s48720 => (case ((reg_name_matches_prefix s48720 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s48730)) => (case ((string_drop s48720 s48730)) of s2 => Some (s, rd, rs1, rs2, s2) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s48580 :: " string " \ \\val _s4840_ : string -> maybe ((bool * bool * bool * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4840 :: \ string \(bool*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4840 s48420 = ( (case ((mul_mnemonic_matches_prefix s48420)) of Some (((high, signed1, signed2), s48430)) => (case ((string_drop s48420 s48430)) of s48440 => (case ((spc_matches_prefix0 s48440)) of Some ((_, s48450)) => (case ((string_drop s48440 s48450)) of s48460 => (case ((reg_name_matches_prefix s48460 :: (( 5 Word.word * ii)) option)) of Some ((rd, s48470)) => (case ((string_drop s48460 s48470)) of s48480 => (case ((sep_matches_prefix s48480)) of Some ((_, s48490)) => (case ((string_drop s48480 s48490)) of s48500 => (case ((reg_name_matches_prefix s48500 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s48510)) => (case ((string_drop s48500 s48510)) of s48520 => (case ((sep_matches_prefix s48520)) of Some ((_, s48530)) => (case ((string_drop s48520 s48530)) of s48540 => (case ((reg_name_matches_prefix s48540 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s48550)) => (case ((string_drop s48540 s48550)) of s1 => Some (high, signed1, signed2, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s48420 :: " string " \ \\val _s4828_ : string -> maybe ((mword ty5 * mword ty5 * string))\\ definition s4828 :: \ string \((5)Word.word*(5)Word.word*string)option \ where \ s4828 s48290 = ( (let s48300 = s48290 in if ((string_startswith s48300 (''c.add''))) then (case ((string_drop s48300 ((string_length (''c.add''))))) of s48310 => (case ((spc_matches_prefix0 s48310)) of Some ((_, s48320)) => (case ((string_drop s48310 s48320)) of s48330 => (case ((reg_name_matches_prefix s48330 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s48340)) => (case ((string_drop s48330 s48340)) of s48350 => (case ((sep_matches_prefix s48350)) of Some ((_, s48360)) => (case ((string_drop s48350 s48360)) of s48370 => (case ((reg_name_matches_prefix s48370 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s48380)) => (case ((string_drop s48370 s48380)) of s1 => Some (rsd, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s48290 :: " string " \ \\val _s4824_ : string -> maybe string\\ definition s4824 :: \ string \(string)option \ where \ s4824 s48250 = ( (let s48260 = s48250 in if ((string_startswith s48260 (''c.ebreak''))) then (case ((string_drop s48260 ((string_length (''c.ebreak''))))) of s1 => Some s1 ) else None))\ for s48250 :: " string " \ \\val _s4812_ : string -> maybe ((mword ty5 * mword ty5 * string))\\ definition s4812 :: \ string \((5)Word.word*(5)Word.word*string)option \ where \ s4812 s48130 = ( (let s48140 = s48130 in if ((string_startswith s48140 (''c.mv''))) then (case ((string_drop s48140 ((string_length (''c.mv''))))) of s48150 => (case ((spc_matches_prefix0 s48150)) of Some ((_, s48160)) => (case ((string_drop s48150 s48160)) of s48170 => (case ((reg_name_matches_prefix s48170 :: (( 5 Word.word * ii)) option)) of Some ((rd, s48180)) => (case ((string_drop s48170 s48180)) of s48190 => (case ((sep_matches_prefix s48190)) of Some ((_, s48200)) => (case ((string_drop s48190 s48200)) of s48210 => (case ((reg_name_matches_prefix s48210 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s48220)) => (case ((string_drop s48210 s48220)) of s1 => Some (rd, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s48130 :: " string " \ \\val _s4804_ : string -> maybe ((mword ty5 * string))\\ definition s4804 :: \ string \((5)Word.word*string)option \ where \ s4804 s48050 = ( (let s48060 = s48050 in if ((string_startswith s48060 (''c.jalr''))) then (case ((string_drop s48060 ((string_length (''c.jalr''))))) of s48070 => (case ((spc_matches_prefix0 s48070)) of Some ((_, s48080)) => (case ((string_drop s48070 s48080)) of s48090 => (case ((reg_name_matches_prefix s48090 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s48100)) => (case ((string_drop s48090 s48100)) of s1 => Some (rs1, s1) ) | _ => None ) ) | _ => None ) ) else None))\ for s48050 :: " string " \ \\val _s4796_ : string -> maybe ((mword ty5 * string))\\ definition s4796 :: \ string \((5)Word.word*string)option \ where \ s4796 s47970 = ( (let s47980 = s47970 in if ((string_startswith s47980 (''c.jr''))) then (case ((string_drop s47980 ((string_length (''c.jr''))))) of s47990 => (case ((spc_matches_prefix0 s47990)) of Some ((_, s48000)) => (case ((string_drop s47990 s48000)) of s48010 => (case ((reg_name_matches_prefix s48010 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s48020)) => (case ((string_drop s48010 s48020)) of s1 => Some (rs1, s1) ) | _ => None ) ) | _ => None ) ) else None))\ for s47970 :: " string " \ \\val _s4784_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4784 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4784 s47850 = ( (let s47860 = s47850 in if ((string_startswith s47860 (''c.sdsp''))) then (case ((string_drop s47860 ((string_length (''c.sdsp''))))) of s47870 => (case ((spc_matches_prefix0 s47870)) of Some ((_, s47880)) => (case ((string_drop s47870 s47880)) of s47890 => (case ((reg_name_matches_prefix s47890 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s47900)) => (case ((string_drop s47890 s47900)) of s47910 => (case ((sep_matches_prefix s47910)) of Some ((_, s47920)) => (case ((string_drop s47910 s47920)) of s47930 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s47930 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s47940)) => (case ((string_drop s47930 s47940)) of s1 => Some (rs2, uimm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s47850 :: " string " \ \\val _s4772_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4772 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4772 s47730 = ( (let s47740 = s47730 in if ((string_startswith s47740 (''c.swsp''))) then (case ((string_drop s47740 ((string_length (''c.swsp''))))) of s47750 => (case ((spc_matches_prefix0 s47750)) of Some ((_, s47760)) => (case ((string_drop s47750 s47760)) of s47770 => (case ((reg_name_matches_prefix s47770 :: (( 5 Word.word * ii)) option)) of Some ((rd, s47780)) => (case ((string_drop s47770 s47780)) of s47790 => (case ((sep_matches_prefix s47790)) of Some ((_, s47800)) => (case ((string_drop s47790 s47800)) of s47810 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s47810 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s47820)) => (case ((string_drop s47810 s47820)) of s1 => Some (rd, uimm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s47730 :: " string " \ \\val _s4760_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4760 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4760 s47610 = ( (let s47620 = s47610 in if ((string_startswith s47620 (''c.ldsp''))) then (case ((string_drop s47620 ((string_length (''c.ldsp''))))) of s47630 => (case ((spc_matches_prefix0 s47630)) of Some ((_, s47640)) => (case ((string_drop s47630 s47640)) of s47650 => (case ((reg_name_matches_prefix s47650 :: (( 5 Word.word * ii)) option)) of Some ((rd, s47660)) => (case ((string_drop s47650 s47660)) of s47670 => (case ((sep_matches_prefix s47670)) of Some ((_, s47680)) => (case ((string_drop s47670 s47680)) of s47690 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s47690 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s47700)) => (case ((string_drop s47690 s47700)) of s1 => Some (rd, uimm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s47610 :: " string " \ \\val _s4748_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4748 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4748 s47490 = ( (let s47500 = s47490 in if ((string_startswith s47500 (''c.lwsp''))) then (case ((string_drop s47500 ((string_length (''c.lwsp''))))) of s47510 => (case ((spc_matches_prefix0 s47510)) of Some ((_, s47520)) => (case ((string_drop s47510 s47520)) of s47530 => (case ((reg_name_matches_prefix s47530 :: (( 5 Word.word * ii)) option)) of Some ((rd, s47540)) => (case ((string_drop s47530 s47540)) of s47550 => (case ((sep_matches_prefix s47550)) of Some ((_, s47560)) => (case ((string_drop s47550 s47560)) of s47570 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s47570 :: (( 6 Word.word * ii)) option)) of Some ((uimm, s47580)) => (case ((string_drop s47570 s47580)) of s1 => Some (rd, uimm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s47490 :: " string " \ \\val _s4736_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4736 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4736 s47370 = ( (let s47380 = s47370 in if ((string_startswith s47380 (''c.slli''))) then (case ((string_drop s47380 ((string_length (''c.slli''))))) of s47390 => (case ((spc_matches_prefix0 s47390)) of Some ((_, s47400)) => (case ((string_drop s47390 s47400)) of s47410 => (case ((reg_name_matches_prefix s47410 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s47420)) => (case ((string_drop s47410 s47420)) of s47430 => (case ((sep_matches_prefix s47430)) of Some ((_, s47440)) => (case ((string_drop s47430 s47440)) of s47450 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s47450 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s47460)) => (case ((string_drop s47450 s47460)) of s1 => Some (rsd, shamt, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s47370 :: " string " \ \\val _s4724_ : string -> maybe ((mword ty3 * mword ty8 * string))\\ definition s4724 :: \ string \((3)Word.word*(8)Word.word*string)option \ where \ s4724 s47250 = ( (let s47260 = s47250 in if ((string_startswith s47260 (''c.bnez''))) then (case ((string_drop s47260 ((string_length (''c.bnez''))))) of s47270 => (case ((spc_matches_prefix0 s47270)) of Some ((_, s47280)) => (case ((string_drop s47270 s47280)) of s47290 => (case ((creg_name_matches_prefix s47290 :: (( 3 Word.word * ii)) option)) of Some ((rs, s47300)) => (case ((string_drop s47290 s47300)) of s47310 => (case ((sep_matches_prefix s47310)) of Some ((_, s47320)) => (case ((string_drop s47310 s47320)) of s47330 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s47330 :: (( 8 Word.word * ii)) option)) of Some ((imm, s47340)) => (case ((string_drop s47330 s47340)) of s1 => Some (rs, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s47250 :: " string " \ \\val _s4712_ : string -> maybe ((mword ty3 * mword ty8 * string))\\ definition s4712 :: \ string \((3)Word.word*(8)Word.word*string)option \ where \ s4712 s47130 = ( (let s47140 = s47130 in if ((string_startswith s47140 (''c.beqz''))) then (case ((string_drop s47140 ((string_length (''c.beqz''))))) of s47150 => (case ((spc_matches_prefix0 s47150)) of Some ((_, s47160)) => (case ((string_drop s47150 s47160)) of s47170 => (case ((creg_name_matches_prefix s47170 :: (( 3 Word.word * ii)) option)) of Some ((rs, s47180)) => (case ((string_drop s47170 s47180)) of s47190 => (case ((sep_matches_prefix s47190)) of Some ((_, s47200)) => (case ((string_drop s47190 s47200)) of s47210 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s47210 :: (( 8 Word.word * ii)) option)) of Some ((imm, s47220)) => (case ((string_drop s47210 s47220)) of s1 => Some (rs, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s47130 :: " string " \ \\val _s4704_ : string -> maybe ((mword ty11 * string))\\ definition s4704 :: \ string \((11)Word.word*string)option \ where \ s4704 s47050 = ( (let s47060 = s47050 in if ((string_startswith s47060 (''c.j''))) then (case ((string_drop s47060 ((string_length (''c.j''))))) of s47070 => (case ((spc_matches_prefix0 s47070)) of Some ((_, s47080)) => (case ((string_drop s47070 s47080)) of s47090 => (case ((hex_bits_11_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s47090 :: (( 11 Word.word * ii)) option)) of Some ((imm, s47100)) => (case ((string_drop s47090 s47100)) of s1 => Some (imm, s1) ) | _ => None ) ) | _ => None ) ) else None))\ for s47050 :: " string " \ \\val _s4692_ : string -> maybe ((mword ty3 * mword ty3 * string))\\ definition s4692 :: \ string \((3)Word.word*(3)Word.word*string)option \ where \ s4692 s46930 = ( (let s46940 = s46930 in if ((string_startswith s46940 (''c.addw''))) then (case ((string_drop s46940 ((string_length (''c.addw''))))) of s46950 => (case ((spc_matches_prefix0 s46950)) of Some ((_, s46960)) => (case ((string_drop s46950 s46960)) of s46970 => (case ((creg_name_matches_prefix s46970 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46980)) => (case ((string_drop s46970 s46980)) of s46990 => (case ((sep_matches_prefix s46990)) of Some ((_, s47000)) => (case ((string_drop s46990 s47000)) of s47010 => (case ((creg_name_matches_prefix s47010 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s47020)) => (case ((string_drop s47010 s47020)) of s1 => Some (rsd, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s46930 :: " string " \ \\val _s4680_ : string -> maybe ((mword ty3 * mword ty3 * string))\\ definition s4680 :: \ string \((3)Word.word*(3)Word.word*string)option \ where \ s4680 s46810 = ( (let s46820 = s46810 in if ((string_startswith s46820 (''c.subw''))) then (case ((string_drop s46820 ((string_length (''c.subw''))))) of s46830 => (case ((spc_matches_prefix0 s46830)) of Some ((_, s46840)) => (case ((string_drop s46830 s46840)) of s46850 => (case ((creg_name_matches_prefix s46850 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46860)) => (case ((string_drop s46850 s46860)) of s46870 => (case ((sep_matches_prefix s46870)) of Some ((_, s46880)) => (case ((string_drop s46870 s46880)) of s46890 => (case ((creg_name_matches_prefix s46890 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s46900)) => (case ((string_drop s46890 s46900)) of s1 => Some (rsd, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s46810 :: " string " \ \\val _s4668_ : string -> maybe ((mword ty3 * mword ty3 * string))\\ definition s4668 :: \ string \((3)Word.word*(3)Word.word*string)option \ where \ s4668 s46690 = ( (let s46700 = s46690 in if ((string_startswith s46700 (''c.and''))) then (case ((string_drop s46700 ((string_length (''c.and''))))) of s46710 => (case ((spc_matches_prefix0 s46710)) of Some ((_, s46720)) => (case ((string_drop s46710 s46720)) of s46730 => (case ((creg_name_matches_prefix s46730 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46740)) => (case ((string_drop s46730 s46740)) of s46750 => (case ((sep_matches_prefix s46750)) of Some ((_, s46760)) => (case ((string_drop s46750 s46760)) of s46770 => (case ((creg_name_matches_prefix s46770 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s46780)) => (case ((string_drop s46770 s46780)) of s1 => Some (rsd, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s46690 :: " string " \ \\val _s4656_ : string -> maybe ((mword ty3 * mword ty3 * string))\\ definition s4656 :: \ string \((3)Word.word*(3)Word.word*string)option \ where \ s4656 s46570 = ( (let s46580 = s46570 in if ((string_startswith s46580 (''c.or''))) then (case ((string_drop s46580 ((string_length (''c.or''))))) of s46590 => (case ((spc_matches_prefix0 s46590)) of Some ((_, s46600)) => (case ((string_drop s46590 s46600)) of s46610 => (case ((creg_name_matches_prefix s46610 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46620)) => (case ((string_drop s46610 s46620)) of s46630 => (case ((sep_matches_prefix s46630)) of Some ((_, s46640)) => (case ((string_drop s46630 s46640)) of s46650 => (case ((creg_name_matches_prefix s46650 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s46660)) => (case ((string_drop s46650 s46660)) of s1 => Some (rsd, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s46570 :: " string " \ \\val _s4644_ : string -> maybe ((mword ty3 * mword ty3 * string))\\ definition s4644 :: \ string \((3)Word.word*(3)Word.word*string)option \ where \ s4644 s46450 = ( (let s46460 = s46450 in if ((string_startswith s46460 (''c.xor''))) then (case ((string_drop s46460 ((string_length (''c.xor''))))) of s46470 => (case ((spc_matches_prefix0 s46470)) of Some ((_, s46480)) => (case ((string_drop s46470 s46480)) of s46490 => (case ((creg_name_matches_prefix s46490 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46500)) => (case ((string_drop s46490 s46500)) of s46510 => (case ((sep_matches_prefix s46510)) of Some ((_, s46520)) => (case ((string_drop s46510 s46520)) of s46530 => (case ((creg_name_matches_prefix s46530 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s46540)) => (case ((string_drop s46530 s46540)) of s1 => Some (rsd, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s46450 :: " string " \ \\val _s4632_ : string -> maybe ((mword ty3 * mword ty3 * string))\\ definition s4632 :: \ string \((3)Word.word*(3)Word.word*string)option \ where \ s4632 s46330 = ( (let s46340 = s46330 in if ((string_startswith s46340 (''c.sub''))) then (case ((string_drop s46340 ((string_length (''c.sub''))))) of s46350 => (case ((spc_matches_prefix0 s46350)) of Some ((_, s46360)) => (case ((string_drop s46350 s46360)) of s46370 => (case ((creg_name_matches_prefix s46370 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46380)) => (case ((string_drop s46370 s46380)) of s46390 => (case ((sep_matches_prefix s46390)) of Some ((_, s46400)) => (case ((string_drop s46390 s46400)) of s46410 => (case ((creg_name_matches_prefix s46410 :: (( 3 Word.word * ii)) option)) of Some ((rs2, s46420)) => (case ((string_drop s46410 s46420)) of s1 => Some (rsd, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s46330 :: " string " \ \\val _s4620_ : string -> maybe ((mword ty3 * mword ty6 * string))\\ definition s4620 :: \ string \((3)Word.word*(6)Word.word*string)option \ where \ s4620 s46210 = ( (let s46220 = s46210 in if ((string_startswith s46220 (''c.andi''))) then (case ((string_drop s46220 ((string_length (''c.andi''))))) of s46230 => (case ((spc_matches_prefix0 s46230)) of Some ((_, s46240)) => (case ((string_drop s46230 s46240)) of s46250 => (case ((creg_name_matches_prefix s46250 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46260)) => (case ((string_drop s46250 s46260)) of s46270 => (case ((sep_matches_prefix s46270)) of Some ((_, s46280)) => (case ((string_drop s46270 s46280)) of s46290 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s46290 :: (( 6 Word.word * ii)) option)) of Some ((imm, s46300)) => (case ((string_drop s46290 s46300)) of s1 => Some (rsd, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s46210 :: " string " \ \\val _s4608_ : string -> maybe ((mword ty3 * mword ty6 * string))\\ definition s4608 :: \ string \((3)Word.word*(6)Word.word*string)option \ where \ s4608 s46090 = ( (let s46100 = s46090 in if ((string_startswith s46100 (''c.srai''))) then (case ((string_drop s46100 ((string_length (''c.srai''))))) of s46110 => (case ((spc_matches_prefix0 s46110)) of Some ((_, s46120)) => (case ((string_drop s46110 s46120)) of s46130 => (case ((creg_name_matches_prefix s46130 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46140)) => (case ((string_drop s46130 s46140)) of s46150 => (case ((sep_matches_prefix s46150)) of Some ((_, s46160)) => (case ((string_drop s46150 s46160)) of s46170 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s46170 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s46180)) => (case ((string_drop s46170 s46180)) of s1 => Some (rsd, shamt, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s46090 :: " string " \ \\val _s4596_ : string -> maybe ((mword ty3 * mword ty6 * string))\\ definition s4596 :: \ string \((3)Word.word*(6)Word.word*string)option \ where \ s4596 s45970 = ( (let s45980 = s45970 in if ((string_startswith s45980 (''c.srli''))) then (case ((string_drop s45980 ((string_length (''c.srli''))))) of s45990 => (case ((spc_matches_prefix0 s45990)) of Some ((_, s46000)) => (case ((string_drop s45990 s46000)) of s46010 => (case ((creg_name_matches_prefix s46010 :: (( 3 Word.word * ii)) option)) of Some ((rsd, s46020)) => (case ((string_drop s46010 s46020)) of s46030 => (case ((sep_matches_prefix s46030)) of Some ((_, s46040)) => (case ((string_drop s46030 s46040)) of s46050 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s46050 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s46060)) => (case ((string_drop s46050 s46060)) of s1 => Some (rsd, shamt, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s45970 :: " string " \ \\val _s4584_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4584 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4584 s45850 = ( (let s45860 = s45850 in if ((string_startswith s45860 (''c.lui''))) then (case ((string_drop s45860 ((string_length (''c.lui''))))) of s45870 => (case ((spc_matches_prefix0 s45870)) of Some ((_, s45880)) => (case ((string_drop s45870 s45880)) of s45890 => (case ((reg_name_matches_prefix s45890 :: (( 5 Word.word * ii)) option)) of Some ((rd, s45900)) => (case ((string_drop s45890 s45900)) of s45910 => (case ((sep_matches_prefix s45910)) of Some ((_, s45920)) => (case ((string_drop s45910 s45920)) of s45930 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s45930 :: (( 6 Word.word * ii)) option)) of Some ((imm, s45940)) => (case ((string_drop s45930 s45940)) of s1 => Some (rd, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s45850 :: " string " \ \\val _s4576_ : string -> maybe ((mword ty6 * string))\\ definition s4576 :: \ string \((6)Word.word*string)option \ where \ s4576 s45770 = ( (let s45780 = s45770 in if ((string_startswith s45780 (''c.addi16sp''))) then (case ((string_drop s45780 ((string_length (''c.addi16sp''))))) of s45790 => (case ((spc_matches_prefix0 s45790)) of Some ((_, s45800)) => (case ((string_drop s45790 s45800)) of s45810 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s45810 :: (( 6 Word.word * ii)) option)) of Some ((imm, s45820)) => (case ((string_drop s45810 s45820)) of s1 => Some (imm, s1) ) | _ => None ) ) | _ => None ) ) else None))\ for s45770 :: " string " \ \\val _s4564_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4564 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4564 s45650 = ( (let s45660 = s45650 in if ((string_startswith s45660 (''c.li''))) then (case ((string_drop s45660 ((string_length (''c.li''))))) of s45670 => (case ((spc_matches_prefix0 s45670)) of Some ((_, s45680)) => (case ((string_drop s45670 s45680)) of s45690 => (case ((reg_name_matches_prefix s45690 :: (( 5 Word.word * ii)) option)) of Some ((rd, s45700)) => (case ((string_drop s45690 s45700)) of s45710 => (case ((sep_matches_prefix s45710)) of Some ((_, s45720)) => (case ((string_drop s45710 s45720)) of s45730 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s45730 :: (( 6 Word.word * ii)) option)) of Some ((imm, s45740)) => (case ((string_drop s45730 s45740)) of s1 => Some (rd, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s45650 :: " string " \ \\val _s4552_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4552 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4552 s45530 = ( (let s45540 = s45530 in if ((string_startswith s45540 (''c.addiw''))) then (case ((string_drop s45540 ((string_length (''c.addiw''))))) of s45550 => (case ((spc_matches_prefix0 s45550)) of Some ((_, s45560)) => (case ((string_drop s45550 s45560)) of s45570 => (case ((reg_name_matches_prefix s45570 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s45580)) => (case ((string_drop s45570 s45580)) of s45590 => (case ((sep_matches_prefix s45590)) of Some ((_, s45600)) => (case ((string_drop s45590 s45600)) of s45610 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s45610 :: (( 6 Word.word * ii)) option)) of Some ((imm, s45620)) => (case ((string_drop s45610 s45620)) of s1 => Some (rsd, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s45530 :: " string " \ \\val _s4544_ : string -> maybe ((mword ty11 * string))\\ definition s4544 :: \ string \((11)Word.word*string)option \ where \ s4544 s45450 = ( (let s45460 = s45450 in if ((string_startswith s45460 (''c.jal''))) then (case ((string_drop s45460 ((string_length (''c.jal''))))) of s45470 => (case ((spc_matches_prefix0 s45470)) of Some ((_, s45480)) => (case ((string_drop s45470 s45480)) of s45490 => (case ((hex_bits_12_matches_prefix0 s45490 :: (( 12 Word.word * ii)) option)) of Some ((v__1216, s45500)) => if (((((subrange_vec_dec v__1216 (( 0 :: int):: ii) (( 0 :: int):: ii) :: 1 Word.word)) = ( 0b0 :: 1 Word.word)))) then (let (imm :: 11 Word.word) = ((subrange_vec_dec v__1216 (( 11 :: int):: ii) (( 1 :: int):: ii) :: 11 Word.word)) in (let (imm :: 11 Word.word) = ((subrange_vec_dec v__1216 (( 11 :: int):: ii) (( 1 :: int):: ii) :: 11 Word.word)) in (case ((string_drop s45490 s45500)) of s1 => Some (imm, s1) ))) else None | _ => None ) ) | _ => None ) ) else None))\ for s45450 :: " string " \ \\val _s4532_ : string -> maybe ((mword ty5 * mword ty6 * string))\\ definition s4532 :: \ string \((5)Word.word*(6)Word.word*string)option \ where \ s4532 s45330 = ( (let s45340 = s45330 in if ((string_startswith s45340 (''c.addi''))) then (case ((string_drop s45340 ((string_length (''c.addi''))))) of s45350 => (case ((spc_matches_prefix0 s45350)) of Some ((_, s45360)) => (case ((string_drop s45350 s45360)) of s45370 => (case ((reg_name_matches_prefix s45370 :: (( 5 Word.word * ii)) option)) of Some ((rsd, s45380)) => (case ((string_drop s45370 s45380)) of s45390 => (case ((sep_matches_prefix s45390)) of Some ((_, s45400)) => (case ((string_drop s45390 s45400)) of s45410 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s45410 :: (( 6 Word.word * ii)) option)) of Some ((nzi, s45420)) => (case ((string_drop s45410 s45420)) of s1 => Some (rsd, nzi, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s45330 :: " string " \ \\val _s4516_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5 * string))\\ definition s4516 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word*string)option \ where \ s4516 s45170 = ( (let s45180 = s45170 in if ((string_startswith s45180 (''c.sd''))) then (case ((string_drop s45180 ((string_length (''c.sd''))))) of s45190 => (case ((spc_matches_prefix0 s45190)) of Some ((_, s45200)) => (case ((string_drop s45190 s45200)) of s45210 => (case ((creg_name_matches_prefix s45210 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s45220)) => (case ((string_drop s45210 s45220)) of s45230 => (case ((sep_matches_prefix s45230)) of Some ((_, s45240)) => (case ((string_drop s45230 s45240)) of s45250 => (case ((creg_name_matches_prefix s45250 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s45260)) => (case ((string_drop s45250 s45260)) of s45270 => (case ((sep_matches_prefix s45270)) of Some ((_, s45280)) => (case ((string_drop s45270 s45280)) of s45290 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s45290 :: (( 8 Word.word * ii)) option)) of Some ((v__1218, s45300)) => if (((((subrange_vec_dec v__1218 (( 2 :: int):: ii) (( 0 :: int):: ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1218 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1218 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (case ((string_drop s45290 s45300)) of s1 => Some (rsc1, rsc2, uimm, s1) ))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s45170 :: " string " \ \\val _s4500_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5 * string))\\ definition s4500 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word*string)option \ where \ s4500 s45010 = ( (let s45020 = s45010 in if ((string_startswith s45020 (''c.sw''))) then (case ((string_drop s45020 ((string_length (''c.sw''))))) of s45030 => (case ((spc_matches_prefix0 s45030)) of Some ((_, s45040)) => (case ((string_drop s45030 s45040)) of s45050 => (case ((creg_name_matches_prefix s45050 :: (( 3 Word.word * ii)) option)) of Some ((rsc1, s45060)) => (case ((string_drop s45050 s45060)) of s45070 => (case ((sep_matches_prefix s45070)) of Some ((_, s45080)) => (case ((string_drop s45070 s45080)) of s45090 => (case ((creg_name_matches_prefix s45090 :: (( 3 Word.word * ii)) option)) of Some ((rsc2, s45100)) => (case ((string_drop s45090 s45100)) of s45110 => (case ((sep_matches_prefix s45110)) of Some ((_, s45120)) => (case ((string_drop s45110 s45120)) of s45130 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s45130 :: (( 7 Word.word * ii)) option)) of Some ((v__1220, s45140)) => if (((((subrange_vec_dec v__1220 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1220 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1220 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (case ((string_drop s45130 s45140)) of s1 => Some (rsc1, rsc2, uimm, s1) ))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s45010 :: " string " \ \\val _s4484_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5 * string))\\ definition s4484 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word*string)option \ where \ s4484 s44850 = ( (let s44860 = s44850 in if ((string_startswith s44860 (''c.ld''))) then (case ((string_drop s44860 ((string_length (''c.ld''))))) of s44870 => (case ((spc_matches_prefix0 s44870)) of Some ((_, s44880)) => (case ((string_drop s44870 s44880)) of s44890 => (case ((creg_name_matches_prefix s44890 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s44900)) => (case ((string_drop s44890 s44900)) of s44910 => (case ((sep_matches_prefix s44910)) of Some ((_, s44920)) => (case ((string_drop s44910 s44920)) of s44930 => (case ((creg_name_matches_prefix s44930 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s44940)) => (case ((string_drop s44930 s44940)) of s44950 => (case ((sep_matches_prefix s44950)) of Some ((_, s44960)) => (case ((string_drop s44950 s44960)) of s44970 => (case ((hex_bits_8_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s44970 :: (( 8 Word.word * ii)) option)) of Some ((v__1222, s44980)) => if (((((subrange_vec_dec v__1222 (( 2 :: int):: ii) (( 0 :: int):: ii) :: 3 Word.word)) = ( 0b000 :: 3 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1222 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1222 (( 7 :: int):: ii) (( 3 :: int):: ii) :: 5 Word.word)) in (case ((string_drop s44970 s44980)) of s1 => Some (rdc, rsc, uimm, s1) ))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s44850 :: " string " \ \\val _s4468_ : string -> maybe ((mword ty3 * mword ty3 * mword ty5 * string))\\ definition s4468 :: \ string \((3)Word.word*(3)Word.word*(5)Word.word*string)option \ where \ s4468 s44690 = ( (let s44700 = s44690 in if ((string_startswith s44700 (''c.lw''))) then (case ((string_drop s44700 ((string_length (''c.lw''))))) of s44710 => (case ((spc_matches_prefix0 s44710)) of Some ((_, s44720)) => (case ((string_drop s44710 s44720)) of s44730 => (case ((creg_name_matches_prefix s44730 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s44740)) => (case ((string_drop s44730 s44740)) of s44750 => (case ((sep_matches_prefix s44750)) of Some ((_, s44760)) => (case ((string_drop s44750 s44760)) of s44770 => (case ((creg_name_matches_prefix s44770 :: (( 3 Word.word * ii)) option)) of Some ((rsc, s44780)) => (case ((string_drop s44770 s44780)) of s44790 => (case ((sep_matches_prefix s44790)) of Some ((_, s44800)) => (case ((string_drop s44790 s44800)) of s44810 => (case ((hex_bits_7_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s44810 :: (( 7 Word.word * ii)) option)) of Some ((v__1224, s44820)) => if (((((subrange_vec_dec v__1224 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1224 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (let (uimm :: 5 Word.word) = ((subrange_vec_dec v__1224 (( 6 :: int):: ii) (( 2 :: int):: ii) :: 5 Word.word)) in (case ((string_drop s44810 s44820)) of s1 => Some (rdc, rsc, uimm, s1) ))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s44690 :: " string " \ \\val _s4456_ : string -> maybe ((mword ty3 * mword ty8 * string))\\ definition s4456 :: \ string \((3)Word.word*(8)Word.word*string)option \ where \ s4456 s44570 = ( (let s44580 = s44570 in if ((string_startswith s44580 (''c.addi4spn''))) then (case ((string_drop s44580 ((string_length (''c.addi4spn''))))) of s44590 => (case ((spc_matches_prefix0 s44590)) of Some ((_, s44600)) => (case ((string_drop s44590 s44600)) of s44610 => (case ((creg_name_matches_prefix s44610 :: (( 3 Word.word * ii)) option)) of Some ((rdc, s44620)) => (case ((string_drop s44610 s44620)) of s44630 => (case ((sep_matches_prefix s44630)) of Some ((_, s44640)) => (case ((string_drop s44630 s44640)) of s44650 => (case ((hex_bits_10_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s44650 :: (( 10 Word.word * ii)) option)) of Some ((v__1226, s44660)) => if (((((subrange_vec_dec v__1226 (( 1 :: int):: ii) (( 0 :: int):: ii) :: 2 Word.word)) = ( 0b00 :: 2 Word.word)))) then (let (nzimm :: 8 Word.word) = ((subrange_vec_dec v__1226 (( 9 :: int):: ii) (( 2 :: int):: ii) :: 8 Word.word)) in (let (nzimm :: 8 Word.word) = ((subrange_vec_dec v__1226 (( 9 :: int):: ii) (( 2 :: int):: ii) :: 8 Word.word)) in (case ((string_drop s44650 s44660)) of s1 => Some (rdc, nzimm, s1) ))) else None | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s44570 :: " string " \ \\val _s4452_ : string -> maybe string\\ definition s4452 :: \ string \(string)option \ where \ s4452 s44530 = ( (let s44540 = s44530 in if ((string_startswith s44540 (''c.nop''))) then (case ((string_drop s44540 ((string_length (''c.nop''))))) of s1 => Some s1 ) else None))\ for s44530 :: " string " \ \\val _s4426_ : string -> maybe ((amoop * word_width * bool * bool * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4426 :: \ string \(amoop*word_width*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4426 s44280 = ( (case ((amo_mnemonic_matches_prefix s44280)) of Some ((op1, s44290)) => (let s44300 = (string_drop s44280 s44290) in if ((string_startswith s44300 (''.''))) then (case ((string_drop s44300 ((string_length (''.''))))) of s44310 => (case ((size_mnemonic_matches_prefix s44310)) of Some ((width, s44320)) => (case ((string_drop s44310 s44320)) of s44330 => (case ((maybe_aq_matches_prefix s44330)) of Some ((aq, s44340)) => (case ((string_drop s44330 s44340)) of s44350 => (case ((maybe_rl_matches_prefix s44350)) of Some ((rl, s44360)) => (case ((string_drop s44350 s44360)) of s44370 => (case ((spc_matches_prefix0 s44370)) of Some ((_, s44380)) => (case ((string_drop s44370 s44380)) of s44390 => (case ((reg_name_matches_prefix s44390 :: (( 5 Word.word * ii)) option)) of Some ((rd, s44400)) => (case ((string_drop s44390 s44400)) of s44410 => (case ((sep_matches_prefix s44410)) of Some ((_, s44420)) => (case ((string_drop s44410 s44420)) of s44430 => (case ((reg_name_matches_prefix s44430 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s44440)) => (case ((string_drop s44430 s44440)) of s44450 => (case ((sep_matches_prefix s44450)) of Some ((_, s44460)) => (let s44470 = (string_drop s44450 s44460) in if ((string_startswith s44470 (''(''))) then (case ((string_drop s44470 ((string_length (''(''))))) of s44480 => (case ((reg_name_matches_prefix s44480 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s44490)) => (let s44500 = (string_drop s44480 s44490) in if ((string_startswith s44500 ('')''))) then (case ((string_drop s44500 ((string_length ('')''))))) of s1 => Some (op1, width, aq, rl, rd, rs2, rs1, s1) ) else None) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ))\ for s44280 :: " string " \ \\val _s4404_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4404 :: \ string \(word_width*bool*bool*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4404 s44050 = ( (let s44060 = s44050 in if ((string_startswith s44060 (''sc.''))) then (case ((string_drop s44060 ((string_length (''sc.''))))) of s44070 => (case ((size_mnemonic_matches_prefix s44070)) of Some ((size1, s44080)) => (case ((string_drop s44070 s44080)) of s44090 => (case ((maybe_aq_matches_prefix s44090)) of Some ((aq, s44100)) => (case ((string_drop s44090 s44100)) of s44110 => (case ((maybe_rl_matches_prefix s44110)) of Some ((rl, s44120)) => (case ((string_drop s44110 s44120)) of s44130 => (case ((spc_matches_prefix0 s44130)) of Some ((_, s44140)) => (case ((string_drop s44130 s44140)) of s44150 => (case ((reg_name_matches_prefix s44150 :: (( 5 Word.word * ii)) option)) of Some ((rd, s44160)) => (case ((string_drop s44150 s44160)) of s44170 => (case ((sep_matches_prefix s44170)) of Some ((_, s44180)) => (case ((string_drop s44170 s44180)) of s44190 => (case ((reg_name_matches_prefix s44190 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s44200)) => (case ((string_drop s44190 s44200)) of s44210 => (case ((sep_matches_prefix s44210)) of Some ((_, s44220)) => (case ((string_drop s44210 s44220)) of s44230 => (case ((reg_name_matches_prefix s44230 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s44240)) => (case ((string_drop s44230 s44240)) of s1 => Some (size1, aq, rl, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s44050 :: " string " \ \\val _s4386_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty5 * string))\\ definition s4386 :: \ string \(word_width*bool*bool*(5)Word.word*(5)Word.word*string)option \ where \ s4386 s43870 = ( (let s43880 = s43870 in if ((string_startswith s43880 (''lr.''))) then (case ((string_drop s43880 ((string_length (''lr.''))))) of s43890 => (case ((size_mnemonic_matches_prefix s43890)) of Some ((size1, s43900)) => (case ((string_drop s43890 s43900)) of s43910 => (case ((maybe_aq_matches_prefix s43910)) of Some ((aq, s43920)) => (case ((string_drop s43910 s43920)) of s43930 => (case ((maybe_rl_matches_prefix s43930)) of Some ((rl, s43940)) => (case ((string_drop s43930 s43940)) of s43950 => (case ((spc_matches_prefix0 s43950)) of Some ((_, s43960)) => (case ((string_drop s43950 s43960)) of s43970 => (case ((reg_name_matches_prefix s43970 :: (( 5 Word.word * ii)) option)) of Some ((rd, s43980)) => (case ((string_drop s43970 s43980)) of s43990 => (case ((sep_matches_prefix s43990)) of Some ((_, s44000)) => (case ((string_drop s43990 s44000)) of s44010 => (case ((reg_name_matches_prefix s44010 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s44020)) => (case ((string_drop s44010 s44020)) of s1 => Some (size1, aq, rl, rd, rs1, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s43870 :: " string " \ \\val _s4374_ : string -> maybe ((mword ty5 * mword ty5 * string))\\ definition s4374 :: \ string \((5)Word.word*(5)Word.word*string)option \ where \ s4374 s43750 = ( (let s43760 = s43750 in if ((string_startswith s43760 (''sfence.vma''))) then (case ((string_drop s43760 ((string_length (''sfence.vma''))))) of s43770 => (case ((spc_matches_prefix0 s43770)) of Some ((_, s43780)) => (case ((string_drop s43770 s43780)) of s43790 => (case ((reg_name_matches_prefix s43790 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s43800)) => (case ((string_drop s43790 s43800)) of s43810 => (case ((sep_matches_prefix s43810)) of Some ((_, s43820)) => (case ((string_drop s43810 s43820)) of s43830 => (case ((reg_name_matches_prefix s43830 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s43840)) => (case ((string_drop s43830 s43840)) of s1 => Some (rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s43750 :: " string " \ \\val _s4370_ : string -> maybe string\\ definition s4370 :: \ string \(string)option \ where \ s4370 s43710 = ( (let s43720 = s43710 in if ((string_startswith s43720 (''wfi''))) then (case ((string_drop s43720 ((string_length (''wfi''))))) of s1 => Some s1 ) else None))\ for s43710 :: " string " \ \\val _s4366_ : string -> maybe string\\ definition s4366 :: \ string \(string)option \ where \ s4366 s43670 = ( (let s43680 = s43670 in if ((string_startswith s43680 (''ebreak''))) then (case ((string_drop s43680 ((string_length (''ebreak''))))) of s1 => Some s1 ) else None))\ for s43670 :: " string " \ \\val _s4362_ : string -> maybe string\\ definition s4362 :: \ string \(string)option \ where \ s4362 s43630 = ( (let s43640 = s43630 in if ((string_startswith s43640 (''sret''))) then (case ((string_drop s43640 ((string_length (''sret''))))) of s1 => Some s1 ) else None))\ for s43630 :: " string " \ \\val _s4358_ : string -> maybe string\\ definition s4358 :: \ string \(string)option \ where \ s4358 s43590 = ( (let s43600 = s43590 in if ((string_startswith s43600 (''mret''))) then (case ((string_drop s43600 ((string_length (''mret''))))) of s1 => Some s1 ) else None))\ for s43590 :: " string " \ \\val _s4354_ : string -> maybe string\\ definition s4354 :: \ string \(string)option \ where \ s4354 s43550 = ( (let s43560 = s43550 in if ((string_startswith s43560 (''ecall''))) then (case ((string_drop s43560 ((string_length (''ecall''))))) of s1 => Some s1 ) else None))\ for s43550 :: " string " \ \\val _s4350_ : string -> maybe string\\ definition s4350 :: \ string \(string)option \ where \ s4350 s43510 = ( (let s43520 = s43510 in if ((string_startswith s43520 (''fence.i''))) then (case ((string_drop s43520 ((string_length (''fence.i''))))) of s1 => Some s1 ) else None))\ for s43510 :: " string " \ \\val _s4338_ : string -> maybe ((mword ty4 * mword ty4 * string))\\ definition s4338 :: \ string \((4)Word.word*(4)Word.word*string)option \ where \ s4338 s43390 = ( (let s43400 = s43390 in if ((string_startswith s43400 (''fence.tso''))) then (case ((string_drop s43400 ((string_length (''fence.tso''))))) of s43410 => (case ((spc_matches_prefix0 s43410)) of Some ((_, s43420)) => (case ((string_drop s43410 s43420)) of s43430 => (case ((fence_bits_matches_prefix s43430 :: (( 4 Word.word * ii)) option)) of Some ((pred, s43440)) => (case ((string_drop s43430 s43440)) of s43450 => (case ((sep_matches_prefix s43450)) of Some ((_, s43460)) => (case ((string_drop s43450 s43460)) of s43470 => (case ((fence_bits_matches_prefix s43470 :: (( 4 Word.word * ii)) option)) of Some ((succ, s43480)) => (case ((string_drop s43470 s43480)) of s1 => Some (pred, succ, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s43390 :: " string " \ \\val _s4326_ : string -> maybe ((mword ty4 * mword ty4 * string))\\ definition s4326 :: \ string \((4)Word.word*(4)Word.word*string)option \ where \ s4326 s43270 = ( (let s43280 = s43270 in if ((string_startswith s43280 (''fence''))) then (case ((string_drop s43280 ((string_length (''fence''))))) of s43290 => (case ((spc_matches_prefix0 s43290)) of Some ((_, s43300)) => (case ((string_drop s43290 s43300)) of s43310 => (case ((fence_bits_matches_prefix s43310 :: (( 4 Word.word * ii)) option)) of Some ((pred, s43320)) => (case ((string_drop s43310 s43320)) of s43330 => (case ((sep_matches_prefix s43330)) of Some ((_, s43340)) => (case ((string_drop s43330 s43340)) of s43350 => (case ((fence_bits_matches_prefix s43350 :: (( 4 Word.word * ii)) option)) of Some ((succ, s43360)) => (case ((string_drop s43350 s43360)) of s1 => Some (pred, succ, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s43270 :: " string " \ \\val _s4309_ : string -> maybe ((sopw * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4309 :: \ string \(sopw*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4309 s43110 = ( (case ((shiftiwop_mnemonic_matches_prefix s43110)) of Some ((op1, s43120)) => (case ((string_drop s43110 s43120)) of s43130 => (case ((spc_matches_prefix0 s43130)) of Some ((_, s43140)) => (case ((string_drop s43130 s43140)) of s43150 => (case ((reg_name_matches_prefix s43150 :: (( 5 Word.word * ii)) option)) of Some ((rd, s43160)) => (case ((string_drop s43150 s43160)) of s43170 => (case ((sep_matches_prefix s43170)) of Some ((_, s43180)) => (case ((string_drop s43170 s43180)) of s43190 => (case ((reg_name_matches_prefix s43190 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s43200)) => (case ((string_drop s43190 s43200)) of s43210 => (case ((sep_matches_prefix s43210)) of Some ((_, s43220)) => (case ((string_drop s43210 s43220)) of s43230 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s43230 :: (( 5 Word.word * ii)) option)) of Some ((shamt, s43240)) => (case ((string_drop s43230 s43240)) of s1 => Some (op1, rd, rs1, shamt, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s43110 :: " string " \ \\val _s4292_ : string -> maybe ((ropw * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4292 :: \ string \(ropw*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4292 s42940 = ( (case ((rtypew_mnemonic_matches_prefix s42940)) of Some ((op1, s42950)) => (case ((string_drop s42940 s42950)) of s42960 => (case ((spc_matches_prefix0 s42960)) of Some ((_, s42970)) => (case ((string_drop s42960 s42970)) of s42980 => (case ((reg_name_matches_prefix s42980 :: (( 5 Word.word * ii)) option)) of Some ((rd, s42990)) => (case ((string_drop s42980 s42990)) of s43000 => (case ((sep_matches_prefix s43000)) of Some ((_, s43010)) => (case ((string_drop s43000 s43010)) of s43020 => (case ((reg_name_matches_prefix s43020 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s43030)) => (case ((string_drop s43020 s43030)) of s43040 => (case ((sep_matches_prefix s43040)) of Some ((_, s43050)) => (case ((string_drop s43040 s43050)) of s43060 => (case ((reg_name_matches_prefix s43060 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s43070)) => (case ((string_drop s43060 s43070)) of s1 => Some (op1, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s42940 :: " string " \ \\val _s4275_ : string -> maybe ((sop * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4275 :: \ string \(sop*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4275 s42770 = ( (case ((shiftw_mnemonic_matches_prefix s42770)) of Some ((op1, s42780)) => (case ((string_drop s42770 s42780)) of s42790 => (case ((spc_matches_prefix0 s42790)) of Some ((_, s42800)) => (case ((string_drop s42790 s42800)) of s42810 => (case ((reg_name_matches_prefix s42810 :: (( 5 Word.word * ii)) option)) of Some ((rd, s42820)) => (case ((string_drop s42810 s42820)) of s42830 => (case ((sep_matches_prefix s42830)) of Some ((_, s42840)) => (case ((string_drop s42830 s42840)) of s42850 => (case ((reg_name_matches_prefix s42850 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s42860)) => (case ((string_drop s42850 s42860)) of s42870 => (case ((sep_matches_prefix s42870)) of Some ((_, s42880)) => (case ((string_drop s42870 s42880)) of s42890 => (case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s42890 :: (( 5 Word.word * ii)) option)) of Some ((shamt, s42900)) => (case ((string_drop s42890 s42900)) of s1 => Some (op1, rd, rs1, shamt, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s42770 :: " string " \ \\val _s4259_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12 * string))\\ definition s4259 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word*string)option \ where \ s4259 s42600 = ( (let s42610 = s42600 in if ((string_startswith s42610 (''addiw''))) then (case ((string_drop s42610 ((string_length (''addiw''))))) of s42620 => (case ((spc_matches_prefix0 s42620)) of Some ((_, s42630)) => (case ((string_drop s42620 s42630)) of s42640 => (case ((reg_name_matches_prefix s42640 :: (( 5 Word.word * ii)) option)) of Some ((rd, s42650)) => (case ((string_drop s42640 s42650)) of s42660 => (case ((sep_matches_prefix s42660)) of Some ((_, s42670)) => (case ((string_drop s42660 s42670)) of s42680 => (case ((reg_name_matches_prefix s42680 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s42690)) => (case ((string_drop s42680 s42690)) of s42700 => (case ((sep_matches_prefix s42700)) of Some ((_, s42710)) => (case ((string_drop s42700 s42710)) of s42720 => (case ((hex_bits_12_matches_prefix0 s42720 :: (( 12 Word.word * ii)) option)) of Some ((imm, s42730)) => (case ((string_drop s42720 s42730)) of s1 => Some (rd, rs1, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s42600 :: " string " \ \\val _s4231_ : string -> maybe ((word_width * bool * bool * mword ty5 * mword ty12 * mword ty5 * string))\\ definition s4231 :: \ string \(word_width*bool*bool*(5)Word.word*(12)Word.word*(5)Word.word*string)option \ where \ s4231 s42320 = ( (let s42330 = s42320 in if ((string_startswith s42330 (''s''))) then (case ((string_drop s42330 ((string_length (''s''))))) of s42340 => (case ((size_mnemonic_matches_prefix s42340)) of Some ((size1, s42350)) => (case ((string_drop s42340 s42350)) of s42360 => (case ((maybe_aq_matches_prefix s42360)) of Some ((aq, s42370)) => (case ((string_drop s42360 s42370)) of s42380 => (case ((maybe_rl_matches_prefix s42380)) of Some ((rl, s42390)) => (case ((string_drop s42380 s42390)) of s42400 => (case ((spc_matches_prefix0 s42400)) of Some ((_, s42410)) => (case ((string_drop s42400 s42410)) of s42420 => (case ((reg_name_matches_prefix s42420 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s42430)) => (case ((string_drop s42420 s42430)) of s42440 => (case ((sep_matches_prefix s42440)) of Some ((_, s42450)) => (case ((string_drop s42440 s42450)) of s42460 => (case ((hex_bits_12_matches_prefix0 s42460 :: (( 12 Word.word * ii)) option)) of Some ((imm, s42470)) => (case ((string_drop s42460 s42470)) of s42480 => (case ((opt_spc_matches_prefix0 s42480)) of Some ((_, s42490)) => (let s42500 = (string_drop s42480 s42490) in if ((string_startswith s42500 (''(''))) then (case ((string_drop s42500 ((string_length (''(''))))) of s42510 => (case ((opt_spc_matches_prefix0 s42510)) of Some ((_, s42520)) => (case ((string_drop s42510 s42520)) of s42530 => (case ((reg_name_matches_prefix s42530 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s42540)) => (case ((string_drop s42530 s42540)) of s42550 => (case ((opt_spc_matches_prefix0 s42550)) of Some ((_, s42560)) => (let s42570 = (string_drop s42550 s42560) in if ((string_startswith s42570 ('')''))) then (case ((string_drop s42570 ((string_length ('')''))))) of s1 => Some (size1, aq, rl, rs2, imm, rs1, s1) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s42320 :: " string " \ \\val _s4201_ : string -> maybe ((word_width * bool * bool * bool * mword ty5 * mword ty12 * mword ty5 * string))\\ definition s4201 :: \ string \(word_width*bool*bool*bool*(5)Word.word*(12)Word.word*(5)Word.word*string)option \ where \ s4201 s42020 = ( (let s42030 = s42020 in if ((string_startswith s42030 (''l''))) then (case ((string_drop s42030 ((string_length (''l''))))) of s42040 => (case ((size_mnemonic_matches_prefix s42040)) of Some ((size1, s42050)) => (case ((string_drop s42040 s42050)) of s42060 => (case ((maybe_u_matches_prefix s42060)) of Some ((is_unsigned, s42070)) => (case ((string_drop s42060 s42070)) of s42080 => (case ((maybe_aq_matches_prefix s42080)) of Some ((aq, s42090)) => (case ((string_drop s42080 s42090)) of s42100 => (case ((maybe_rl_matches_prefix s42100)) of Some ((rl, s42110)) => (case ((string_drop s42100 s42110)) of s42120 => (case ((spc_matches_prefix0 s42120)) of Some ((_, s42130)) => (case ((string_drop s42120 s42130)) of s42140 => (case ((reg_name_matches_prefix s42140 :: (( 5 Word.word * ii)) option)) of Some ((rd, s42150)) => (case ((string_drop s42140 s42150)) of s42160 => (case ((sep_matches_prefix s42160)) of Some ((_, s42170)) => (case ((string_drop s42160 s42170)) of s42180 => (case ((hex_bits_12_matches_prefix0 s42180 :: (( 12 Word.word * ii)) option)) of Some ((imm, s42190)) => (case ((string_drop s42180 s42190)) of s42200 => (case ((opt_spc_matches_prefix0 s42200)) of Some ((_, s42210)) => (let s42220 = (string_drop s42200 s42210) in if ((string_startswith s42220 (''(''))) then (case ((string_drop s42220 ((string_length (''(''))))) of s42230 => (case ((opt_spc_matches_prefix0 s42230)) of Some ((_, s42240)) => (case ((string_drop s42230 s42240)) of s42250 => (case ((reg_name_matches_prefix s42250 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s42260)) => (case ((string_drop s42250 s42260)) of s42270 => (case ((opt_spc_matches_prefix0 s42270)) of Some ((_, s42280)) => (let s42290 = (string_drop s42270 s42280) in if ((string_startswith s42290 ('')''))) then (case ((string_drop s42290 ((string_length ('')''))))) of s1 => Some (size1, is_unsigned, aq, rl, rd, imm, rs1, s1) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s42020 :: " string " \ \\val _s4184_ : string -> maybe ((rop * mword ty5 * mword ty5 * mword ty5 * string))\\ definition s4184 :: \ string \(rop*(5)Word.word*(5)Word.word*(5)Word.word*string)option \ where \ s4184 s41860 = ( (case ((rtype_mnemonic_matches_prefix s41860)) of Some ((op1, s41870)) => (case ((string_drop s41860 s41870)) of s41880 => (case ((spc_matches_prefix0 s41880)) of Some ((_, s41890)) => (case ((string_drop s41880 s41890)) of s41900 => (case ((reg_name_matches_prefix s41900 :: (( 5 Word.word * ii)) option)) of Some ((rd, s41910)) => (case ((string_drop s41900 s41910)) of s41920 => (case ((sep_matches_prefix s41920)) of Some ((_, s41930)) => (case ((string_drop s41920 s41930)) of s41940 => (case ((reg_name_matches_prefix s41940 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s41950)) => (case ((string_drop s41940 s41950)) of s41960 => (case ((sep_matches_prefix s41960)) of Some ((_, s41970)) => (case ((string_drop s41960 s41970)) of s41980 => (case ((reg_name_matches_prefix s41980 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s41990)) => (case ((string_drop s41980 s41990)) of s1 => Some (op1, rd, rs1, rs2, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s41860 :: " string " \ \\val _s4167_ : string -> maybe ((sop * mword ty5 * mword ty5 * mword ty6 * string))\\ definition s4167 :: \ string \(sop*(5)Word.word*(5)Word.word*(6)Word.word*string)option \ where \ s4167 s41690 = ( (case ((shiftiop_mnemonic_matches_prefix s41690)) of Some ((op1, s41700)) => (case ((string_drop s41690 s41700)) of s41710 => (case ((spc_matches_prefix0 s41710)) of Some ((_, s41720)) => (case ((string_drop s41710 s41720)) of s41730 => (case ((reg_name_matches_prefix s41730 :: (( 5 Word.word * ii)) option)) of Some ((rd, s41740)) => (case ((string_drop s41730 s41740)) of s41750 => (case ((sep_matches_prefix s41750)) of Some ((_, s41760)) => (case ((string_drop s41750 s41760)) of s41770 => (case ((reg_name_matches_prefix s41770 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s41780)) => (case ((string_drop s41770 s41780)) of s41790 => (case ((sep_matches_prefix s41790)) of Some ((_, s41800)) => (case ((string_drop s41790 s41800)) of s41810 => (case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s41810 :: (( 6 Word.word * ii)) option)) of Some ((shamt, s41820)) => (case ((string_drop s41810 s41820)) of s1 => Some (op1, rd, rs1, shamt, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s41690 :: " string " \ \\val _s4150_ : string -> maybe ((iop * mword ty5 * mword ty5 * mword ty12 * string))\\ definition s4150 :: \ string \(iop*(5)Word.word*(5)Word.word*(12)Word.word*string)option \ where \ s4150 s41520 = ( (case ((itype_mnemonic_matches_prefix s41520)) of Some ((op1, s41530)) => (case ((string_drop s41520 s41530)) of s41540 => (case ((spc_matches_prefix0 s41540)) of Some ((_, s41550)) => (case ((string_drop s41540 s41550)) of s41560 => (case ((reg_name_matches_prefix s41560 :: (( 5 Word.word * ii)) option)) of Some ((rd, s41570)) => (case ((string_drop s41560 s41570)) of s41580 => (case ((sep_matches_prefix s41580)) of Some ((_, s41590)) => (case ((string_drop s41580 s41590)) of s41600 => (case ((reg_name_matches_prefix s41600 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s41610)) => (case ((string_drop s41600 s41610)) of s41620 => (case ((sep_matches_prefix s41620)) of Some ((_, s41630)) => (case ((string_drop s41620 s41630)) of s41640 => (case ((hex_bits_12_matches_prefix0 s41640 :: (( 12 Word.word * ii)) option)) of Some ((imm, s41650)) => (case ((string_drop s41640 s41650)) of s1 => Some (op1, rd, rs1, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s41520 :: " string " \ \\val _s4133_ : string -> maybe ((bop * mword ty5 * mword ty5 * mword ty13 * string))\\ definition s4133 :: \ string \(bop*(5)Word.word*(5)Word.word*(13)Word.word*string)option \ where \ s4133 s41350 = ( (case ((btype_mnemonic_matches_prefix s41350)) of Some ((op1, s41360)) => (case ((string_drop s41350 s41360)) of s41370 => (case ((spc_matches_prefix0 s41370)) of Some ((_, s41380)) => (case ((string_drop s41370 s41380)) of s41390 => (case ((reg_name_matches_prefix s41390 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s41400)) => (case ((string_drop s41390 s41400)) of s41410 => (case ((sep_matches_prefix s41410)) of Some ((_, s41420)) => (case ((string_drop s41410 s41420)) of s41430 => (case ((reg_name_matches_prefix s41430 :: (( 5 Word.word * ii)) option)) of Some ((rs2, s41440)) => (case ((string_drop s41430 s41440)) of s41450 => (case ((sep_matches_prefix s41450)) of Some ((_, s41460)) => (case ((string_drop s41450 s41460)) of s41470 => (case ((hex_bits_13_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s41470 :: (( 13 Word.word * ii)) option)) of Some ((imm, s41480)) => (case ((string_drop s41470 s41480)) of s1 => Some (op1, rs1, rs2, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s41350 :: " string " \ \\val _s4117_ : string -> maybe ((mword ty5 * mword ty5 * mword ty12 * string))\\ definition s4117 :: \ string \((5)Word.word*(5)Word.word*(12)Word.word*string)option \ where \ s4117 s41180 = ( (let s41190 = s41180 in if ((string_startswith s41190 (''jalr''))) then (case ((string_drop s41190 ((string_length (''jalr''))))) of s41200 => (case ((spc_matches_prefix0 s41200)) of Some ((_, s41210)) => (case ((string_drop s41200 s41210)) of s41220 => (case ((reg_name_matches_prefix s41220 :: (( 5 Word.word * ii)) option)) of Some ((rd, s41230)) => (case ((string_drop s41220 s41230)) of s41240 => (case ((sep_matches_prefix s41240)) of Some ((_, s41250)) => (case ((string_drop s41240 s41250)) of s41260 => (case ((reg_name_matches_prefix s41260 :: (( 5 Word.word * ii)) option)) of Some ((rs1, s41270)) => (case ((string_drop s41260 s41270)) of s41280 => (case ((sep_matches_prefix s41280)) of Some ((_, s41290)) => (case ((string_drop s41280 s41290)) of s41300 => (case ((hex_bits_12_matches_prefix0 s41300 :: (( 12 Word.word * ii)) option)) of Some ((imm, s41310)) => (case ((string_drop s41300 s41310)) of s1 => Some (rd, rs1, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s41180 :: " string " \ \\val _s4105_ : string -> maybe ((mword ty5 * mword ty21 * string))\\ definition s4105 :: \ string \((5)Word.word*(21)Word.word*string)option \ where \ s4105 s41060 = ( (let s41070 = s41060 in if ((string_startswith s41070 (''jal''))) then (case ((string_drop s41070 ((string_length (''jal''))))) of s41080 => (case ((spc_matches_prefix0 s41080)) of Some ((_, s41090)) => (case ((string_drop s41080 s41090)) of s41100 => (case ((reg_name_matches_prefix s41100 :: (( 5 Word.word * ii)) option)) of Some ((rd, s41110)) => (case ((string_drop s41100 s41110)) of s41120 => (case ((sep_matches_prefix s41120)) of Some ((_, s41130)) => (case ((string_drop s41120 s41130)) of s41140 => (case ((hex_bits_21_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s41140 :: (( 21 Word.word * ii)) option)) of Some ((imm, s41150)) => (case ((string_drop s41140 s41150)) of s1 => Some (rd, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) else None))\ for s41060 :: " string " \ \\val _s4092_ : string -> maybe ((uop * mword ty5 * mword ty20 * string))\\ definition s4092 :: \ string \(uop*(5)Word.word*(20)Word.word*string)option \ where \ s4092 s40940 = ( (case ((utype_mnemonic_matches_prefix s40940)) of Some ((op1, s40950)) => (case ((string_drop s40940 s40950)) of s40960 => (case ((spc_matches_prefix0 s40960)) of Some ((_, s40970)) => (case ((string_drop s40960 s40970)) of s40980 => (case ((reg_name_matches_prefix s40980 :: (( 5 Word.word * ii)) option)) of Some ((rd, s40990)) => (case ((string_drop s40980 s40990)) of s41000 => (case ((sep_matches_prefix s41000)) of Some ((_, s41010)) => (case ((string_drop s41000 s41010)) of s41020 => (case ((hex_bits_20_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict s41020 :: (( 20 Word.word * ii)) option)) of Some ((imm, s41030)) => (case ((string_drop s41020 s41030)) of s1 => Some (op1, rd, imm, s1) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ) ) | _ => None ))\ for s40940 :: " string " definition assembly_matches_prefix :: \ string \(ast*int)option \ where \ assembly_matches_prefix arg1 = ( (let s41040 = arg1 in if ((case ((s4092 s41040 :: ((uop * 5 Word.word * 20 Word.word * string))option)) of Some ((op1, rd, imm, s1)) => True | _ => False )) then (case (s4092 s41040 :: (( uop * 5 Word.word * 20 Word.word * string)) option) of (Some ((op1, rd, imm, s1))) => Some (UTYPE (imm, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4105 s41040 :: (( 5 Word.word * 21 Word.word * string))option)) of Some ((rd, imm, s1)) => True | _ => False )) then (case (s4105 s41040 :: (( 5 Word.word * 21 Word.word * string)) option) of (Some ((rd, imm, s1))) => Some (RISCV_JAL (imm, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4117 s41040 :: (( 5 Word.word * 5 Word.word * 12 Word.word * string))option)) of Some ((rd, rs1, imm, s1)) => True | _ => False )) then (case (s4117 s41040 :: (( 5 Word.word * 5 Word.word * 12 Word.word * string)) option) of (Some ((rd, rs1, imm, s1))) => Some (RISCV_JALR (imm, rs1, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4133 s41040 :: ((bop * 5 Word.word * 5 Word.word * 13 Word.word * string))option)) of Some ((op1, rs1, rs2, imm, s1)) => True | _ => False )) then (case (s4133 s41040 :: (( bop * 5 Word.word * 5 Word.word * 13 Word.word * string)) option) of (Some ((op1, rs1, rs2, imm, s1))) => Some (BTYPE (imm, rs2, rs1, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4150 s41040 :: ((iop * 5 Word.word * 5 Word.word * 12 Word.word * string))option)) of Some ((op1, rd, rs1, imm, s1)) => True | _ => False )) then (case (s4150 s41040 :: (( iop * 5 Word.word * 5 Word.word * 12 Word.word * string)) option) of (Some ((op1, rd, rs1, imm, s1))) => Some (ITYPE (imm, rs1, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4167 s41040 :: ((sop * 5 Word.word * 5 Word.word * 6 Word.word * string))option)) of Some ((op1, rd, rs1, shamt, s1)) => True | _ => False )) then (case (s4167 s41040 :: (( sop * 5 Word.word * 5 Word.word * 6 Word.word * string)) option) of (Some ((op1, rd, rs1, shamt, s1))) => Some (SHIFTIOP (shamt, rs1, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4184 s41040 :: ((rop * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((op1, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s4184 s41040 :: (( rop * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((op1, rd, rs1, rs2, s1))) => Some (RTYPE (rs2, rs1, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4201 s41040 :: ((word_width * bool * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word * string))option)) of Some ((size1, is_unsigned, aq, rl, rd, imm, rs1, s1)) => True | _ => False )) then (case (s4201 s41040 :: (( word_width * bool * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word * string)) option) of (Some ((size1, is_unsigned, aq, rl, rd, imm, rs1, s1))) => Some (LOAD (imm, rs1, rd, is_unsigned, size1, aq, rl), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4231 s41040 :: ((word_width * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word * string))option)) of Some ((size1, aq, rl, rs2, imm, rs1, s1)) => True | _ => False )) then (case (s4231 s41040 :: (( word_width * bool * bool * 5 Word.word * 12 Word.word * 5 Word.word * string)) option) of (Some ((size1, aq, rl, rs2, imm, rs1, s1))) => Some (STORE (imm, rs2, rs1, size1, aq, rl), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4259 s41040 :: (( 5 Word.word * 5 Word.word * 12 Word.word * string))option)) of Some ((rd, rs1, imm, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4259 s41040 :: (( 5 Word.word * 5 Word.word * 12 Word.word * string)) option) of (Some ((rd, rs1, imm, s1))) => Some (ADDIW (imm, rs1, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4275 s41040 :: ((sop * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((op1, rd, rs1, shamt, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4275 s41040 :: (( sop * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((op1, rd, rs1, shamt, s1))) => Some (SHIFTW (shamt, rs1, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4292 s41040 :: ((ropw * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((op1, rd, rs1, rs2, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4292 s41040 :: (( ropw * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((op1, rd, rs1, rs2, s1))) => Some (RTYPEW (rs2, rs1, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4309 s41040 :: ((sopw * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((op1, rd, rs1, shamt, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4309 s41040 :: (( sopw * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((op1, rd, rs1, shamt, s1))) => Some (SHIFTIWOP (shamt, rs1, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4326 s41040 :: (( 4 Word.word * 4 Word.word * string))option)) of Some ((pred, succ, s1)) => True | _ => False )) then (case (s4326 s41040 :: (( 4 Word.word * 4 Word.word * string)) option) of (Some ((pred, succ, s1))) => Some (FENCE (pred, succ), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4338 s41040 :: (( 4 Word.word * 4 Word.word * string))option)) of Some ((pred, succ, s1)) => True | _ => False )) then (case (s4338 s41040 :: (( 4 Word.word * 4 Word.word * string)) option) of (Some ((pred, succ, s1))) => Some (FENCE_TSO (pred, succ), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4350 s41040)) of Some (s1) => True | _ => False )) then (case s4350 s41040 of (Some (s1)) => Some (FENCEI () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4354 s41040)) of Some (s1) => True | _ => False )) then (case s4354 s41040 of (Some (s1)) => Some (ECALL () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4358 s41040)) of Some (s1) => True | _ => False )) then (case s4358 s41040 of (Some (s1)) => Some (MRET () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4362 s41040)) of Some (s1) => True | _ => False )) then (case s4362 s41040 of (Some (s1)) => Some (SRET () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4366 s41040)) of Some (s1) => True | _ => False )) then (case s4366 s41040 of (Some (s1)) => Some (EBREAK () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4370 s41040)) of Some (s1) => True | _ => False )) then (case s4370 s41040 of (Some (s1)) => Some (WFI () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4374 s41040 :: (( 5 Word.word * 5 Word.word * string))option)) of Some ((rs1, rs2, s1)) => True | _ => False )) then (case (s4374 s41040 :: (( 5 Word.word * 5 Word.word * string)) option) of (Some ((rs1, rs2, s1))) => Some (SFENCE_VMA (rs1, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4386 s41040 :: ((word_width * bool * bool * 5 Word.word * 5 Word.word * string))option)) of Some ((size1, aq, rl, rd, rs1, s1)) => True | _ => False )) then (case (s4386 s41040 :: (( word_width * bool * bool * 5 Word.word * 5 Word.word * string)) option) of (Some ((size1, aq, rl, rd, rs1, s1))) => Some (LOADRES (aq, rl, rs1, size1, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4404 s41040 :: ((word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((size1, aq, rl, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s4404 s41040 :: (( word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((size1, aq, rl, rd, rs1, rs2, s1))) => Some (STORECON (aq, rl, rs2, rs1, size1, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4426 s41040 :: ((amoop * word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((op1, width, aq, rl, rd, rs2, rs1, s1)) => True | _ => False )) then (case (s4426 s41040 :: (( amoop * word_width * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((op1, width, aq, rl, rd, rs2, rs1, s1))) => Some (AMO (op1, aq, rl, rs2, rs1, width, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4452 s41040)) of Some (s1) => True | _ => False )) then (case s4452 s41040 of (Some (s1)) => Some (C_NOP () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4456 s41040 :: (( 3 Word.word * 8 Word.word * string))option)) of Some ((rdc, nzimm, s1)) => (nzimm \ ( 0x00 :: 8 Word.word)) | _ => False )) then (case (s4456 s41040 :: (( 3 Word.word * 8 Word.word * string)) option) of (Some ((rdc, nzimm, s1))) => Some (C_ADDI4SPN (rdc, nzimm), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4468 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string))option)) of Some ((rdc, rsc, uimm, s1)) => True | _ => False )) then (case (s4468 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string)) option) of (Some ((rdc, rsc, uimm, s1))) => Some (C_LW (uimm, rsc, rdc), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4484 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string))option)) of Some ((rdc, rsc, uimm, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4484 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string)) option) of (Some ((rdc, rsc, uimm, s1))) => Some (C_LD (uimm, rsc, rdc), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4500 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string))option)) of Some ((rsc1, rsc2, uimm, s1)) => True | _ => False )) then (case (s4500 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string)) option) of (Some ((rsc1, rsc2, uimm, s1))) => Some (C_SW (uimm, rsc1, rsc2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4516 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string))option)) of Some ((rsc1, rsc2, uimm, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4516 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string)) option) of (Some ((rsc1, rsc2, uimm, s1))) => Some (C_SD (uimm, rsc1, rsc2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4532 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rsd, nzi, s1)) => ((((nzi \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))) | _ => False )) then (case (s4532 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rsd, nzi, s1))) => Some (C_ADDI (nzi, rsd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4544 s41040 :: (( 11 Word.word * string))option)) of Some ((imm, s1)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s4544 s41040 :: (( 11 Word.word * string)) option) of (Some ((imm, s1))) => Some (C_JAL imm, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4552 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rsd, imm, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4552 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rsd, imm, s1))) => Some (C_ADDIW (imm, rsd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4564 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rd, imm, s1)) => (rd \ zreg) | _ => False )) then (case (s4564 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rd, imm, s1))) => Some (C_LI (imm, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4576 s41040 :: (( 6 Word.word * string))option)) of Some ((imm, s1)) => (imm \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s4576 s41040 :: (( 6 Word.word * string)) option) of (Some ((imm, s1))) => Some (C_ADDI16SP imm, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4584 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rd, imm, s1)) => ((((rd \ zreg))) \ ((((((rd \ sp))) \ (((imm \ ( 0b000000 :: 6 Word.word)))))))) | _ => False )) then (case (s4584 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rd, imm, s1))) => Some (C_LUI (imm, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4596 s41040 :: (( 3 Word.word * 6 Word.word * string))option)) of Some ((rsd, shamt, s1)) => (shamt \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s4596 s41040 :: (( 3 Word.word * 6 Word.word * string)) option) of (Some ((rsd, shamt, s1))) => Some (C_SRLI (shamt, rsd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4608 s41040 :: (( 3 Word.word * 6 Word.word * string))option)) of Some ((rsd, shamt, s1)) => (shamt \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s4608 s41040 :: (( 3 Word.word * 6 Word.word * string)) option) of (Some ((rsd, shamt, s1))) => Some (C_SRAI (shamt, rsd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4620 s41040 :: (( 3 Word.word * 6 Word.word * string))option)) of Some ((rsd, imm, s1)) => True | _ => False )) then (case (s4620 s41040 :: (( 3 Word.word * 6 Word.word * string)) option) of (Some ((rsd, imm, s1))) => Some (C_ANDI (imm, rsd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4632 s41040 :: (( 3 Word.word * 3 Word.word * string))option)) of Some ((rsd, rs2, s1)) => True | _ => False )) then (case (s4632 s41040 :: (( 3 Word.word * 3 Word.word * string)) option) of (Some ((rsd, rs2, s1))) => Some (C_SUB (rsd, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4644 s41040 :: (( 3 Word.word * 3 Word.word * string))option)) of Some ((rsd, rs2, s1)) => True | _ => False )) then (case (s4644 s41040 :: (( 3 Word.word * 3 Word.word * string)) option) of (Some ((rsd, rs2, s1))) => Some (C_XOR (rsd, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4656 s41040 :: (( 3 Word.word * 3 Word.word * string))option)) of Some ((rsd, rs2, s1)) => True | _ => False )) then (case (s4656 s41040 :: (( 3 Word.word * 3 Word.word * string)) option) of (Some ((rsd, rs2, s1))) => Some (C_OR (rsd, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4668 s41040 :: (( 3 Word.word * 3 Word.word * string))option)) of Some ((rsd, rs2, s1)) => True | _ => False )) then (case (s4668 s41040 :: (( 3 Word.word * 3 Word.word * string)) option) of (Some ((rsd, rs2, s1))) => Some (C_AND (rsd, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4680 s41040 :: (( 3 Word.word * 3 Word.word * string))option)) of Some ((rsd, rs2, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4680 s41040 :: (( 3 Word.word * 3 Word.word * string)) option) of (Some ((rsd, rs2, s1))) => Some (C_SUBW (rsd, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4692 s41040 :: (( 3 Word.word * 3 Word.word * string))option)) of Some ((rsd, rs2, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4692 s41040 :: (( 3 Word.word * 3 Word.word * string)) option) of (Some ((rsd, rs2, s1))) => Some (C_ADDW (rsd, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4704 s41040 :: (( 11 Word.word * string))option)) of Some ((imm, s1)) => True | _ => False )) then (case (s4704 s41040 :: (( 11 Word.word * string)) option) of (Some ((imm, s1))) => Some (C_J imm, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4712 s41040 :: (( 3 Word.word * 8 Word.word * string))option)) of Some ((rs, imm, s1)) => True | _ => False )) then (case (s4712 s41040 :: (( 3 Word.word * 8 Word.word * string)) option) of (Some ((rs, imm, s1))) => Some (C_BEQZ (imm, rs), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4724 s41040 :: (( 3 Word.word * 8 Word.word * string))option)) of Some ((rs, imm, s1)) => True | _ => False )) then (case (s4724 s41040 :: (( 3 Word.word * 8 Word.word * string)) option) of (Some ((rs, imm, s1))) => Some (C_BNEZ (imm, rs), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4736 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rsd, shamt, s1)) => ((((shamt \ ( 0b000000 :: 6 Word.word)))) \ (((rsd \ zreg)))) | _ => False )) then (case (s4736 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rsd, shamt, s1))) => Some (C_SLLI (shamt, rsd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4748 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rd, uimm, s1)) => (rd \ zreg) | _ => False )) then (case (s4748 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rd, uimm, s1))) => Some (C_LWSP (uimm, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4760 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rd, uimm, s1)) => ((((rd \ zreg))) \ ((((( 32 :: int)::ii) = (( 64 :: int)::ii))))) | _ => False )) then (case (s4760 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rd, uimm, s1))) => Some (C_LDSP (uimm, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4772 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rd, uimm, s1)) => True | _ => False )) then (case (s4772 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rd, uimm, s1))) => Some (C_SWSP (uimm, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4784 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rs2, uimm, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4784 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rs2, uimm, s1))) => Some (C_SDSP (uimm, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4796 s41040 :: (( 5 Word.word * string))option)) of Some ((rs1, s1)) => (rs1 \ zreg) | _ => False )) then (case (s4796 s41040 :: (( 5 Word.word * string)) option) of (Some ((rs1, s1))) => Some (C_JR rs1, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4804 s41040 :: (( 5 Word.word * string))option)) of Some ((rs1, s1)) => (rs1 \ zreg) | _ => False )) then (case (s4804 s41040 :: (( 5 Word.word * string)) option) of (Some ((rs1, s1))) => Some (C_JALR rs1, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4812 s41040 :: (( 5 Word.word * 5 Word.word * string))option)) of Some ((rd, rs2, s1)) => ((((rd \ zreg))) \ (((rs2 \ zreg)))) | _ => False )) then (case (s4812 s41040 :: (( 5 Word.word * 5 Word.word * string)) option) of (Some ((rd, rs2, s1))) => Some (C_MV (rd, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4824 s41040)) of Some (s1) => True | _ => False )) then (case s4824 s41040 of (Some (s1)) => Some (C_EBREAK () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4828 s41040 :: (( 5 Word.word * 5 Word.word * string))option)) of Some ((rsd, rs2, s1)) => ((((rsd \ zreg))) \ (((rs2 \ zreg)))) | _ => False )) then (case (s4828 s41040 :: (( 5 Word.word * 5 Word.word * string)) option) of (Some ((rsd, rs2, s1))) => Some (C_ADD (rsd, rs2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4840 s41040 :: ((bool * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((high, signed1, signed2, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s4840 s41040 :: (( bool * bool * bool * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((high, signed1, signed2, rd, rs1, rs2, s1))) => Some (MUL (rs2, rs1, rd, high, signed1, signed2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4857 s41040 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((s, rd, rs1, rs2, s2)) => True | _ => False )) then (case (s4857 s41040 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((s, rd, rs1, rs2, s2))) => Some (DIV (rs2, rs1, rd, s), ((string_length arg1)) - ((string_length s2))) ) else if ((case ((s4875 s41040 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((s, rd, rs1, rs2, s2)) => True | _ => False )) then (case (s4875 s41040 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((s, rd, rs1, rs2, s2))) => Some (REM (rs2, rs1, rd, s), ((string_length arg1)) - ((string_length s2))) ) else if ((case ((s4893 s41040 :: (( 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((rd, rs1, rs2, s1)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4893 s41040 :: (( 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((rd, rs1, rs2, s1))) => Some (MULW (rs2, rs1, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4909 s41040 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((s, rd, rs1, rs2, s2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4909 s41040 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((s, rd, rs1, rs2, s2))) => Some (DIVW (rs2, rs1, rd, s), ((string_length arg1)) - ((string_length s2))) ) else if ((case ((s4928 s41040 :: ((bool * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((s, rd, rs1, rs2, s2)) => ((( 32 :: int)::ii) = (( 64 :: int)::ii)) | _ => False )) then (case (s4928 s41040 :: (( bool * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((s, rd, rs1, rs2, s2))) => Some (REMW (rs2, rs1, rd, s), ((string_length arg1)) - ((string_length s2))) ) else if ((case ((s4947 s41040 :: ((csrop * 5 Word.word * 12 Word.word * 5 Word.word * string))option)) of Some ((op1, rd, csr, rs1, s1)) => True | _ => False )) then (case (s4947 s41040 :: (( csrop * 5 Word.word * 12 Word.word * 5 Word.word * string)) option) of (Some ((op1, rd, csr, rs1, s1))) => Some (CSR (csr, rs1, rd, True, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4965 s41040 :: ((csrop * 5 Word.word * 12 Word.word * 5 Word.word * string))option)) of Some ((op1, rd, csr, rs1, s1)) => True | _ => False )) then (case (s4965 s41040 :: (( csrop * 5 Word.word * 12 Word.word * 5 Word.word * string)) option) of (Some ((op1, rd, csr, rs1, s1))) => Some (CSR (csr, rs1, rd, False, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4982 s41040)) of Some (s1) => True | _ => False )) then (case s4982 s41040 of (Some (s1)) => Some (URET () , ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4986 s41040 :: (( 6 Word.word * string))option)) of Some ((imm, s1)) => True | _ => False )) then (case (s4986 s41040 :: (( 6 Word.word * string)) option) of (Some ((imm, s1))) => Some (C_NOP_HINT imm, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4992 s41040 :: (( 5 Word.word * string))option)) of Some ((rsd, s1)) => (rsd \ zreg) | _ => False )) then (case (s4992 s41040 :: (( 5 Word.word * string)) option) of (Some ((rsd, s1))) => Some (C_ADDI_HINT rsd, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s4998 s41040 :: (( 6 Word.word * string))option)) of Some ((imm, s1)) => True | _ => False )) then (case (s4998 s41040 :: (( 6 Word.word * string)) option) of (Some ((imm, s1))) => Some (C_LI_HINT imm, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5004 s41040 :: (( 6 Word.word * string))option)) of Some ((imm, s1)) => (imm \ ( 0b000000 :: 6 Word.word)) | _ => False )) then (case (s5004 s41040 :: (( 6 Word.word * string)) option) of (Some ((imm, s1))) => Some (C_LUI_HINT imm, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5010 s41040 :: (( 5 Word.word * string))option)) of Some ((rs2, s1)) => (rs2 \ zreg) | _ => False )) then (case (s5010 s41040 :: (( 5 Word.word * string)) option) of (Some ((rs2, s1))) => Some (C_MV_HINT rs2, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5016 s41040 :: (( 5 Word.word * string))option)) of Some ((rs2, s1)) => (rs2 \ zreg) | _ => False )) then (case (s5016 s41040 :: (( 5 Word.word * string)) option) of (Some ((rs2, s1))) => Some (C_ADD_HINT rs2, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5022 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rsd, shamt, s1)) => ((((shamt = ( 0b000000 :: 6 Word.word)))) \ (((rsd = zreg)))) | _ => False )) then (case (s5022 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rsd, shamt, s1))) => Some (C_SLLI_HINT (shamt, rsd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5031 s41040 :: (( 3 Word.word * string))option)) of Some ((rsd, s1)) => True | _ => False )) then (case (s5031 s41040 :: (( 3 Word.word * string)) option) of (Some ((rsd, s1))) => Some (C_SRLI_HINT rsd, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5037 s41040 :: (( 3 Word.word * string))option)) of Some ((rsd, s1)) => True | _ => False )) then (case (s5037 s41040 :: (( 3 Word.word * string)) option) of (Some ((rsd, s1))) => Some (C_SRAI_HINT rsd, ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5043 s41040 :: (( 4 Word.word * 4 Word.word * 5 Word.word * 5 Word.word * 4 Word.word * string))option)) of Some ((pred, succ, rs, rd, fm, s1)) => (((((((fm \ ( 0x0 :: 4 Word.word)))) \ (((fm \ ( 0x8 :: 4 Word.word))))))) \ ((((((rs \ ( 0b00000 :: 5 Word.word)))) \ (((rd \ ( 0b00000 :: 5 Word.word)))))))) | _ => False )) then (case (s5043 s41040 :: (( 4 Word.word * 4 Word.word * 5 Word.word * 5 Word.word * 4 Word.word * string)) option) of (Some ((pred, succ, rs, rd, fm, s1))) => Some (FENCE_RESERVED (fm, pred, succ, rs, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5061 s41040 :: (( 5 Word.word * 5 Word.word * 12 Word.word * string))option)) of Some ((rd, rs, imm, s1)) => ((((imm \ ( 0x000 :: 12 Word.word)))) \ ((((((rs \ zreg))) \ (((rd \ zreg))))))) | _ => False )) then (case (s5061 s41040 :: (( 5 Word.word * 5 Word.word * 12 Word.word * string)) option) of (Some ((rd, rs, imm, s1))) => Some (FENCEI_RESERVED (imm, rs, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5073 s41040 :: ((word_width * 5 Word.word * 12 Word.word * 5 Word.word * string))option)) of Some ((width, rd, imm, rs1, s1)) => True | _ => False )) then (case (s5073 s41040 :: (( word_width * 5 Word.word * 12 Word.word * 5 Word.word * string)) option) of (Some ((width, rd, imm, rs1, s1))) => Some (LOAD_FP (imm, rs1, rd, width), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5097 s41040 :: ((word_width * 5 Word.word * 12 Word.word * 5 Word.word * string))option)) of Some ((width, rs2, imm, rs1, s1)) => True | _ => False )) then (case (s5097 s41040 :: (( word_width * 5 Word.word * 12 Word.word * 5 Word.word * string)) option) of (Some ((width, rs2, imm, rs1, s1))) => Some (STORE_FP (imm, rs2, rs1, width), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5121 s41040 :: ((f_madd_op_S * 5 Word.word * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((op1, rd, rs1, rs2, rs3, rm, s1)) => True | _ => False )) then (case (s5121 s41040 :: (( f_madd_op_S * 5 Word.word * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((op1, rd, rs1, rs2, rs3, rm, s1))) => Some (F_MADD_TYPE_S (rs3, rs2, rs1, rm, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5146 s41040 :: ((f_bin_rm_op_S * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((op1, rd, rs1, rs2, rm, s1)) => True | _ => False )) then (case (s5146 s41040 :: (( f_bin_rm_op_S * 5 Word.word * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((op1, rd, rs1, rs2, rm, s1))) => Some (F_BIN_RM_TYPE_S (rs2, rs1, rm, rd, op1), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5167 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FSQRT_S, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5167 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FSQRT_S, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FSQRT_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5184 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FCVT_W_S, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5184 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FCVT_W_S, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_W_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5201 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FCVT_WU_S, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5201 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FCVT_WU_S, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_WU_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5218 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FCVT_S_W, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5218 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FCVT_S_W, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_W), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5235 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FCVT_S_WU, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5235 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FCVT_S_WU, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_WU), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5252 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FCVT_L_S, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5252 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FCVT_L_S, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_L_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5269 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FCVT_LU_S, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5269 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FCVT_LU_S, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_LU_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5286 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FCVT_S_L, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5286 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FCVT_S_L, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_L), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5303 s41040 :: ((f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string))option)) of Some ((FCVT_S_LU, rd, rs1, rm, s1)) => True | _ => False )) then (case (s5303 s41040 :: (( f_un_rm_op_S * 5 Word.word * 5 Word.word * rounding_mode * string)) option) of (Some ((FCVT_S_LU, rd, rs1, rm, s1))) => Some (F_UN_RM_TYPE_S (rs1, rm, rd, FCVT_S_LU), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5320 s41040 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((FSGNJ_S, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s5320 s41040 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((FSGNJ_S, rd, rs1, rs2, s1))) => Some (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJ_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5337 s41040 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((FSGNJN_S, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s5337 s41040 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((FSGNJN_S, rd, rs1, rs2, s1))) => Some (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJN_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5354 s41040 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((FSGNJX_S, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s5354 s41040 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((FSGNJX_S, rd, rs1, rs2, s1))) => Some (F_BIN_TYPE_S (rs2, rs1, rd, FSGNJX_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5371 s41040 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((FMIN_S, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s5371 s41040 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((FMIN_S, rd, rs1, rs2, s1))) => Some (F_BIN_TYPE_S (rs2, rs1, rd, FMIN_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5388 s41040 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((FMAX_S, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s5388 s41040 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((FMAX_S, rd, rs1, rs2, s1))) => Some (F_BIN_TYPE_S (rs2, rs1, rd, FMAX_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5405 s41040 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((FEQ_S, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s5405 s41040 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((FEQ_S, rd, rs1, rs2, s1))) => Some (F_BIN_TYPE_S (rs2, rs1, rd, FEQ_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5422 s41040 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((FLT_S, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s5422 s41040 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((FLT_S, rd, rs1, rs2, s1))) => Some (F_BIN_TYPE_S (rs2, rs1, rd, FLT_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5439 s41040 :: ((f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string))option)) of Some ((FLE_S, rd, rs1, rs2, s1)) => True | _ => False )) then (case (s5439 s41040 :: (( f_bin_op_S * 5 Word.word * 5 Word.word * 5 Word.word * string)) option) of (Some ((FLE_S, rd, rs1, rs2, s1))) => Some (F_BIN_TYPE_S (rs2, rs1, rd, FLE_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5456 s41040 :: ((f_un_op_S * 5 Word.word * 5 Word.word * string))option)) of Some ((FMV_X_W, rd, rs1, s1)) => True | _ => False )) then (case (s5456 s41040 :: (( f_un_op_S * 5 Word.word * 5 Word.word * string)) option) of (Some ((FMV_X_W, rd, rs1, s1))) => Some (F_UN_TYPE_S (rs1, rd, FMV_X_W), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5469 s41040 :: ((f_un_op_S * 5 Word.word * 5 Word.word * string))option)) of Some ((FMV_W_X, rd, rs1, s1)) => True | _ => False )) then (case (s5469 s41040 :: (( f_un_op_S * 5 Word.word * 5 Word.word * string)) option) of (Some ((FMV_W_X, rd, rs1, s1))) => Some (F_UN_TYPE_S (rs1, rd, FMV_W_X), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5482 s41040 :: ((f_un_op_S * 5 Word.word * 5 Word.word * string))option)) of Some ((FCLASS_S, rd, rs1, s1)) => True | _ => False )) then (case (s5482 s41040 :: (( f_un_op_S * 5 Word.word * 5 Word.word * string)) option) of (Some ((FCLASS_S, rd, rs1, s1))) => Some (F_UN_TYPE_S (rs1, rd, FCLASS_S), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5495 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rd, imm, s1)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s5495 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rd, imm, s1))) => Some (C_FLWSP (imm, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5507 s41040 :: (( 5 Word.word * 6 Word.word * string))option)) of Some ((rd, uimm, s1)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s5507 s41040 :: (( 5 Word.word * 6 Word.word * string)) option) of (Some ((rd, uimm, s1))) => Some (C_FSWSP (uimm, rd), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5519 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string))option)) of Some ((rdc, rsc, uimm, s1)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s5519 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string)) option) of (Some ((rdc, rsc, uimm, s1))) => Some (C_FLW (uimm, rsc, rdc), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5535 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string))option)) of Some ((rsc1, rsc2, uimm, s1)) => ((( 32 :: int)::ii) = (( 32 :: int)::ii)) | _ => False )) then (case (s5535 s41040 :: (( 3 Word.word * 3 Word.word * 5 Word.word * string)) option) of (Some ((rsc1, rsc2, uimm, s1))) => Some (C_FSW (uimm, rsc1, rsc2), ((string_length arg1)) - ((string_length s1))) ) else if ((case ((s5551 s41040 :: (( 32 Word.word * string))option)) of Some ((s, s2)) => True | _ => False )) then (case (s5551 s41040 :: (( 32 Word.word * string)) option) of (Some ((s, s2))) => Some (ILLEGAL s, ((string_length arg1)) - ((string_length s2))) ) else if ((case ((s5559 s41040 :: (( 16 Word.word * string))option)) of Some ((s, s2)) => True | _ => False )) then (case (s5559 s41040 :: (( 16 Word.word * string)) option) of (Some ((s, s2))) => Some (C_ILLEGAL s, ((string_length arg1)) - ((string_length s2))) ) else None))\ for arg1 :: " string " definition print_insn :: \ ast \((register_value),(string),(exception))monad \ where \ print_insn insn = ( assembly_forwards insn )\ for insn :: " ast " \ \\val decode : mword ty32 -> M ast\\ definition decode :: \(32)Word.word \((register_value),(ast),(exception))monad \ where \ decode bv = ( encdec_backwards bv )\ for bv :: "(32)Word.word " \ \\val decodeCompressed : mword ty16 -> M ast\\ definition decodeCompressed :: \(16)Word.word \((register_value),(ast),(exception))monad \ where \ decodeCompressed bv = ( encdec_compressed_backwards bv )\ for bv :: "(16)Word.word " \ \\val ext_init : unit -> unit\\ definition ext_init :: \ unit \ unit \ where \ ext_init _ = ( () )\ \ \\val ext_fetch_hook : FetchResult -> FetchResult\\ definition ext_fetch_hook :: \ FetchResult \ FetchResult \ where \ ext_fetch_hook f = ( f )\ for f :: " FetchResult " \ \\val ext_pre_step_hook : unit -> unit\\ definition ext_pre_step_hook :: \ unit \ unit \ where \ ext_pre_step_hook _ = ( () )\ \ \\val ext_post_step_hook : unit -> unit\\ definition ext_post_step_hook :: \ unit \ unit \ where \ ext_post_step_hook _ = ( () )\ \ \\val ext_post_decode_hook : ast -> M ast\\ definition ext_post_decode_hook :: \ ast \((register_value),(ast),(exception))monad \ where \ ext_post_decode_hook x = ( return x )\ for x :: " ast " \ \\val isRVC : mword ty16 -> bool\\ definition isRVC :: \(16)Word.word \ bool \ where \ isRVC h = ( \ (((((subrange_vec_dec h (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = ( 0b11 :: 2 Word.word)))))\ for h :: "(16)Word.word " \ \\val fetch : unit -> M FetchResult\\ definition fetch :: \ unit \((register_value),(FetchResult),(exception))monad \ where \ fetch _ = ( (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__0 :: 32 Word.word) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__1 :: 32 Word.word) . (case ((ext_fetch_check_pc w__0 w__1)) of Ext_FetchAddr_Error (e) => return (F_Ext_Error e) | Ext_FetchAddr_OK (use_pc) => or_boolM (return (((((access_vec_dec use_pc (( 0 :: int)::ii))) \ B0)))) (and_boolM (return (((((access_vec_dec use_pc (( 1 :: int)::ii))) \ B0)))) (haveRVC () \ ((\ (w__2 :: bool) . return ((\ w__2)))))) \ ((\ (w__4 :: bool) . if w__4 then (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__5 :: 32 Word.word) . return (F_Error (E_Fetch_Addr_Align () , w__5)))) else (translateAddr use_pc (Execute () ) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__6 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__6 of TR_Failure ((e, _)) => (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__7 :: 32 Word.word) . return (F_Error (e, w__7)))) | TR_Address ((ppclo, _)) => (mem_read (Execute () ) ppclo (( 2 :: int)::ii) False False False :: ( ( 16 Word.word)MemoryOpResult) M) \ ((\ (w__8 :: ( 16 Word.word) MemoryOpResult) . (case w__8 of MemException (e) => (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__9 :: 32 Word.word) . return (F_Error (e, w__9)))) | MemValue (ilo) => if ((isRVC ilo)) then return (F_RVC ilo) else (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__10 :: 32 Word.word) . (let (PC_hi :: xlenbits) = ((add_vec_int w__10 (( 2 :: int)::ii) :: 32 Word.word)) in (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__11 :: 32 Word.word) . (case ((ext_fetch_check_pc w__11 PC_hi)) of Ext_FetchAddr_Error (e) => return (F_Ext_Error e) | Ext_FetchAddr_OK (use_pc_hi) => (translateAddr use_pc_hi (Execute () ) :: ( (( 32 Word.word), ExceptionType)TR_Result) M) \ ((\ (w__12 :: (( 32 Word.word), ExceptionType) TR_Result) . (case w__12 of TR_Failure ((e, _)) => return (F_Error (e, PC_hi)) | TR_Address ((ppchi, _)) => (mem_read (Execute () ) ppchi (( 2 :: int)::ii) False False False :: ( ( 16 Word.word)MemoryOpResult) M) \ ((\ (w__13 :: ( 16 Word.word) MemoryOpResult) . return ((case w__13 of MemException (e) => F_Error (e, PC_hi) | MemValue (ihi) => F_Base ((concat_vec ihi ilo :: 32 Word.word)) )))) ))) )))))) ))) ))))) ))))))\ \ \\val step : ii -> M bool\\ definition step :: \ int \((register_value),(bool),(exception))monad \ where \ step step_no = ( (let (_ :: unit) = (ext_pre_step_hook () ) in (write_reg minstret_written_ref False \ read_reg cur_privilege_ref) \ ((\ (w__0 :: Privilege) . dispatchInterrupt w__0 \ ((\ (w__1 :: ((InterruptType * Privilege))option) . (case w__1 of Some ((intr, priv)) => (let (_ :: unit) = (if ((get_config_print_instr () )) then print_bits0 (''Handling interrupt: '') ((interruptType_to_bits intr :: 8 Word.word)) else () ) in handle_interrupt intr priv \ return (RETIRE_FAIL, False)) | None => fetch () \ ((\ (w__2 :: FetchResult) . (let (f :: FetchResult) = (ext_fetch_hook w__2) in (case f of F_Ext_Error (e) => (let (_ :: unit) = (ext_handle_fetch_check_error e) in return (RETIRE_FAIL, False)) | F_Error ((e, addr)) => handle_mem_exception addr e \ return (RETIRE_FAIL, False) | F_RVC (h) => decodeCompressed h \ ((\ ast . ((if ((get_config_print_instr () )) then read_reg cur_privilege_ref \ ((\ (w__3 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__4 :: 32 Word.word) . print_insn ast \ ((\ (w__5 :: string) . return ((print_dbg (((@) (''['') (((@) ((stringFromInteger step_no)) (((@) (''] ['') (((@) ((privLevel_to_str w__3)) (((@) ('']: '') (((@) ((string_of_bits w__4)) (((@) ('' ('') (((@) ((string_of_bits h)) (((@) ('') '') w__5)))))))))))))))))))))))))) else return () ) \ haveRVC () ) \ ((\ (w__6 :: bool) . if w__6 then (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__7 :: 32 Word.word) . (write_reg nextPC_ref ((add_vec_int w__7 (( 2 :: int)::ii) :: 32 Word.word)) \ ext_post_decode_hook ast) \ ((\ (w__8 :: ast) . execute w__8 \ ((\ (w__9 :: Retired) . return (w__9, True))))))) else handle_illegal () \ return (RETIRE_FAIL, True))))) | F_Base (w) => decode w \ ((\ ast . ((if ((get_config_print_instr () )) then read_reg cur_privilege_ref \ ((\ (w__11 :: Privilege) . (read_reg PC_ref :: ( 32 Word.word) M) \ ((\ (w__12 :: 32 Word.word) . print_insn ast \ ((\ (w__13 :: string) . return ((print_dbg (((@) (''['') (((@) ((stringFromInteger step_no)) (((@) (''] ['') (((@) ((privLevel_to_str w__11)) (((@) ('']: '') (((@) ((string_of_bits w__12)) (((@) ('' ('') (((@) ((string_of_bits w)) (((@) ('') '') w__13)))))))))))))))))))))))))) else return () ) \ (read_reg PC_ref :: ( 32 Word.word) M)) \ ((\ (w__14 :: 32 Word.word) . (write_reg nextPC_ref ((add_vec_int w__14 (( 4 :: int)::ii) :: 32 Word.word)) \ ext_post_decode_hook ast) \ ((\ (w__15 :: ast) . execute w__15 \ ((\ (w__16 :: Retired) . return (w__16, True))))))))) )))) ) \ ((\ varstup . (let ((retired :: Retired), (stepped :: bool)) = varstup in (tick_pc () \ (case retired of RETIRE_SUCCESS => retire_instruction () | RETIRE_FAIL => return () )) \ ((let (_ :: unit) = (ext_post_step_hook () ) in return stepped)))))))))))\ for step_no :: " int " \ \\val loop : unit -> M unit\\ definition loop :: \ unit \((register_value),(unit),(exception))monad \ where \ loop _ = ( (let insns_per_tick = (plat_insns_per_tick () ) in (let (i :: ii) = ((( 0 :: int)::ii)) in (let (step_no :: ii) = ((( 0 :: int)::ii)) in (whileM (i, step_no) ((\ varstup . (let (i, step_no) = varstup in read_reg htif_done_ref \ ((\ (w__0 :: bool) . return ((\ w__0))))))) ((\ varstup . (let (i, step_no) = varstup in step step_no \ ((\ stepped . (let (step_no :: ii) = (if stepped then step_no + (( 1 :: int)::ii) else step_no) in read_reg htif_done_ref \ ((\ (w__1 :: bool) . (if w__1 then (read_reg htif_exit_code_ref :: ( 64 Word.word) M) \ ((\ (w__2 :: 64 Word.word) . (let exit_val = (Word.uint w__2) in return ((let (_ :: unit) = (if (((exit_val = (( 0 :: int)::ii)))) then print_endline (''SUCCESS'') else print_int (''FAILURE: '') exit_val) in i))))) else (let i = (i + (( 1 :: int)::ii)) in if (((i = insns_per_tick))) then (tick_clock () \ tick_platform () ) \ return (( 0 :: int)::ii) else return i)) \ ((\ (i :: ii) . return (i, step_no)))))))))))) \ ((\ varstup . (let ((i :: ii), (step_no :: ii)) = varstup in return () )))))))\ \ \\val init_model : unit -> M unit\\ definition init_model :: \ unit \((register_value),(unit),(exception))monad \ where \ init_model _ = ( ((init_platform () \ init_sys () ) \ init_vmem () ) \ ((let (_ :: unit) = (ext_init () ) in ext_init_regs () )))\ end