diff options
| author | Thomas Bauereiss | 2018-08-29 15:35:44 +0100 |
|---|---|---|
| committer | Thomas Bauereiss | 2018-08-29 15:35:44 +0100 |
| commit | 07e3591e2427db2d9407d554ac57984ca566c6ed (patch) | |
| tree | 9fdd56300b7b0fde4ddecfaed2587bb937cb69ff | |
| parent | 6860b871787df1341c94f4239904fef1743f8625 (diff) | |
Updated snapshots for Isabelle 2018
60 files changed, 9414 insertions, 8316 deletions
diff --git a/lib/isabelle/Makefile b/lib/isabelle/Makefile index 56591140..57459b82 100644 --- a/lib/isabelle/Makefile +++ b/lib/isabelle/Makefile @@ -25,7 +25,7 @@ heap-img: thys $(EXTRA_THYS) ROOT manual: heap-img manual/Manual.thy manual/ROOT manual/document/root.tex cp output/document/session_graph.pdf manual/document/Sail_session_graph.pdf make -C $(RISCV_DIR) Riscv_duopod.thy - isabelle build -d manual Sail-Manual + isabelle build -d manual -d $(RISCV_DIR) Sail-Manual Sail2_instr_kinds.thy: ../../src/lem_interp/sail2_instr_kinds.lem lem -isa -outdir . -auxiliary_level none -lib ../../src/lem_interp -lib ../../src/gen_lib $< diff --git a/lib/isabelle/manual/Manual.thy b/lib/isabelle/manual/Manual.thy index 1ac300b4..f151072a 100644 --- a/lib/isabelle/manual/Manual.thy +++ b/lib/isabelle/manual/Manual.thy @@ -1,10 +1,10 @@ (*<*) theory Manual imports - Sail.Sail2_state_lemmas - Sail.Sail2_operators_mwords_lemmas - Sail.Hoare - Riscv_duopod.Riscv_duopod_lemmas + "Sail.Sail2_state_lemmas" + "Sail.Sail2_operators_mwords_lemmas" + "Sail.Hoare" + "Sail-RISC-V-Duopod.Riscv_duopod_lemmas" begin declare [[show_question_marks = false]] diff --git a/lib/isabelle/manual/ROOT b/lib/isabelle/manual/ROOT index a5923f26..aae60daa 100644 --- a/lib/isabelle/manual/ROOT +++ b/lib/isabelle/manual/ROOT @@ -1,9 +1,4 @@ -session "Riscv_duopod" in "../../../riscv" = "Sail" + - options [document = false] - theories - Riscv_duopod_lemmas - -session "Sail-Manual" = "Riscv_duopod" + +session "Sail-Manual" = "Sail-RISC-V-Duopod" + options [document = pdf, document_output = "output"] theories Manual @@ -2,3 +2,8 @@ session "Sail-RISC-V" = "Sail" + options [document = false] theories Riscv_lemmas + +session "Sail-RISC-V-Duopod" = "Sail" + + options [document = false] + theories + Riscv_duopod_lemmas diff --git a/snapshots/isabelle/Manual.pdf b/snapshots/isabelle/Manual.pdf Binary files differindex 62066f8f..1a33757d 100644 --- a/snapshots/isabelle/Manual.pdf +++ b/snapshots/isabelle/Manual.pdf diff --git a/snapshots/isabelle/Manual.thy b/snapshots/isabelle/Manual.thy index c9eaa18a..f151072a 100644 --- a/snapshots/isabelle/Manual.thy +++ b/snapshots/isabelle/Manual.thy @@ -1,10 +1,10 @@ (*<*) theory Manual imports - Sail.State_lemmas - Sail.Sail_operators_mwords_lemmas - Sail.Hoare - "riscv/Riscv_duopod_lemmas" + "Sail.Sail2_state_lemmas" + "Sail.Sail2_operators_mwords_lemmas" + "Sail.Hoare" + "Sail-RISC-V-Duopod.Riscv_duopod_lemmas" begin declare [[show_question_marks = false]] @@ -60,10 +60,12 @@ This uses the following options: Isabelle definitions can then be generated by passing the @{verbatim "-isa"} flag to Lem. In order for Lem to find the Sail library, the subdirectories @{path "src/gen_lib"} and @{path "src/lem_interp"} of Sail will have to be added to Lem's include path using the -@{verbatim "-lib"} option, e.g. +@{verbatim "-lib"} option. In the following example, these directories are also associated with +the session name @{session Sail}, so that Lem generates import statements with session-qualified +theory names such as @{theory Sail.Hoare}. @{verbatim [display] -"lem -isa -outdir . -lib ../src/lem_interp -lib ../src/gen_lib +"lem -isa -outdir . -lib Sail=../src/lem_interp -lib Sail=../src/gen_lib riscv_extras.lem riscv_duopod_types.lem riscv_duopod.lem"} For further examples, see the @{path Makefile}s of the other specifications included in the Sail @@ -168,8 +170,8 @@ following variants: \<^item> @{term quot_vec_maybe} returns an option type, with @{lemma "quot_vec_maybe w 0 = None" by (auto simp: quot_vec_maybe_def quot_bv_def arith_op_bv_no0_def)}. \<^item> @{term quot_vec_fail} is monadic and either returns the result or raises an exception. - \<^item> @{term quot_vec_oracle} is monadic and uses the @{term Undefined} effect in the exception case - to fill the result with bits drawn from a bitstream oracle. + \<^item> @{term quot_vec_nondet} is monadic and uses the @{term Undefined} effect in the exception case + to fill the result with bits chosen nondeterministically. \<^item> @{term quot_vec} is pure and returns an arbitrary (but fixed) value in the exception case, currently defined as follows: For the bitlist representation of bitvectors, @{term "quot_vec w 0"} returns a list filled with @{term BU}, while for the machine word @@ -183,8 +185,8 @@ Vectors in Sail are mapped to lists in Isabelle, except for bitvectors, which ar Both increasing and decreasing indexing order are supported by having two versions for each operation that involves indexing, such as @{term update_list_inc} and @{term update_list_dec}, or @{term subrange_list_inc} and @{term subrange_list_dec}. These operations are defined in the -theory @{theory Sail_values}, while @{theory Sail_values_lemmas} provides simplification rules -such as +theory @{theory Sail.Sail2_values}, while @{theory Sail.Sail2_values_lemmas} provides simplification +rules such as @{lemma "access_list_inc xs i = xs ! nat i" by auto} \\ @{thm access_list_dec_nth} @@ -209,7 +211,7 @@ into multiple clauses with concrete bitvector lengths. This is not enabled by d so Sail generates Lem definitions using bit lists unless the @{verbatim "-lem_mwords"} command line flag is used. -The theory @{theory Sail_values} defines a (Lem) typeclass @{verbatim Bitvector}, which provides +The theory @{theory Sail.Sail2_values} defines a (Lem) typeclass @{verbatim Bitvector}, which provides an interface to some basic bitvector operations and has instantiations for both bit lists and machine words. It is mainly intended for internal use in the Sail library,\<^footnote>\<open>Lem typeclasses are not very convenient to use in Isabelle, as they get translated to dictionaries that have to be passed to functions @@ -219,7 +221,7 @@ representations. For use in Sail specifications, wrappers are defined in the th right theory is automatically added to the generated files, depending on which bitvector representation is used. Hence, bitvector operations can be referred to in the Sail source code using uniform names, e.g.~@{term add_vec}, @{term update_vec_dec}, or @{term subrange_vec_inc}. -The theory @{theory Sail_operators_mwords_lemmas} sets up simplification rules that relate these +The theory @{theory Sail.Sail2_operators_mwords_lemmas} sets up simplification rules that relate these operations to the native operations in Isabelle, e.g. @{lemma "add_vec l r = l + r" by simp} \\ @@ -235,9 +237,9 @@ for integration with relaxed memory models. For the sequential case, we use a s exceptions and nondeterminism). The generated definitions use the free monad, and the sequential case is supported via a lifting -to the state monad defined in the theory @{theory State}. Simplification rules are set up in the -theory @{theory State_lemmas}, allowing seamless reasoning about the generated definitions in terms -of the state monad.\<close> +to the state monad defined in the theory @{theory Sail.Sail2_state}. Simplification rules are set up +in the theory @{theory Sail.Sail2_state_lemmas}, allowing seamless reasoning about the generated +definitions in terms of the state monad.\<close> subsubsection \<open>State Monad \label{sec:state-monad}\<close> @@ -263,9 +265,6 @@ The @{type sequential_state} record has the following fields: of the last announced memory write, if any. \<^item> The @{term last_exclusive_operation_was_load} flag is used to determine whether exclusive operations can succeed. - \<^item> The function stored in the @{term next_bool} field together with the seed in the @{term seed} - field are used as a random bit generator for undefined values. The @{term next_bool} - function takes the current seed as an argument and returns a @{type bool} and the next seed. The library defines several combinators and wrappers in addition to the standard monadic bind and return (called @{term bindS} and @{term returnS} here, where the suffix @{term S} differentiates them @@ -288,7 +287,7 @@ exception handler as arguments. The exception mechanism is also used to implement early returns by throwing and catching return values: A function body with one or more early returns of type @{typ 'a} (and exception type @{typ 'e}) is lifted to a monadic expression with exception type @{typ "('a + 'e)"} using -@{term liftSR}, such that an early return of the value @{term a} throws @{term "Inl a"}, and a +@{term liftRS}, such that an early return of the value @{term a} throws @{term "Inl a"}, and a regular exception @{term e} is thrown as @{term "Inr e"}. The function body is then wrapped in @{term catch_early_returnS} to lower it back to the default monad and exception type. These liftings and lowerings are automatically inserted by Sail for functions with early returns.\<^footnote>\<open>To be @@ -308,8 +307,8 @@ instances.\<close> subsubsection \<open>Free Monad \label{sec:free-monad}\<close> -text \<open>In addition to the state monad, the theory @{theory Prompt_monad} defines a free monad of an -effect datatype. A monadic expression either returns a pure value @{term a}, denoted +text \<open>In addition to the state monad, the theory @{theory Sail.Sail2_prompt_monad} defines a free +monad of an effect datatype. A monadic expression either returns a pure value @{term a}, denoted @{term "Done a"}, or it has an effect. The latter can be a failure or an exception, or an effect together with a continuation. For example, @{term \<open>Read_reg ''PC'' k\<close>} represents a request to read the register @{term PC} and continue as @{term k}, which is a function that takes the @@ -383,7 +382,7 @@ register state record: Sail aims to generate Isabelle definitions that can be used with either the state or the free monad. To achieve this, the definitions are generated using the free monad, and a lifting to the state monad is provided together with simplification rules. These include generic simplification rules -(proved in the theory @{theory State_lemmas}) such as +(proved in the theory @{theory Sail.Sail2_state_lemmas}) such as @{thm [display] liftState_return[where r = "(get_regval, set_regval)"] liftState_bind[where r = "(get_regval, set_regval)"] @@ -403,7 +402,7 @@ section \<open>Example Proof \label{sec:ex-proof}\<close> text \<open>As a toy example for illustration, we prove that the add instruction in the RISC-V duopod actually performs an addition. We consider the sequential case and use the state monad. The -theory @{theory Hoare} defines (a shallow embedding of) a simple Hoare logic, where +theory @{theory Sail.Hoare} defines (a shallow embedding of) a simple Hoare logic, where @{term "PrePost P f Q"} denotes a triple of a precondition @{term P}, monadic expression @{term f}, and postcondition @{term Q}. Its validity is defined by @{thm [display] PrePost_def} @@ -412,7 +411,9 @@ case, defined as @{thm [display, names_short] PrePostE_def} The theory includes standard proof rules for both of these variants, in particular rules giving weakest preconditions of the predefined primitives of the monad, collected under the names -@{attribute PrePost_intro} and @{attribute PrePostE_intro}, respectively. +@{attribute PrePost_atomI} for atoms such as @{term return} and @{attribute PrePost_compositeI} +for composites such as @{term bind} (or @{attribute PrePostE_atomI} and +@{attribute PrePostE_compositeI}, respectively, for the quadruple variant). The instruction we are considering is defined as @{thm [display] execute_ITYPE.simps[of _ rs for rs]} @@ -448,16 +449,17 @@ lemma shows "PrePostE pre (liftS instr) post (\<lambda>_ _. False)" unfolding pre_def instr_def post_def - by (simp add: rX_def wX_def cong: bindS_cong if_cong split del: if_split) - (rule PrePostE_strengthen_pre, (rule PrePostE_intro)+, auto simp: uint_0_iff) + by (simp add: rX_def wX_def liftState_simp cong: bindS_cong if_cong split del: if_split) + (rule PrePostE_strengthen_pre, (rule PrePostE_atomI PrePostE_compositeI)+, auto simp: uint_0_iff) text \<open>The proof begins with a simplification step, which not only unfolds the definitions of the auxiliary functions @{term rX} and @{term wX}, but also performs the lifting from the free monad to the state monad. We apply the rule @{thm [source] PrePostE_strengthen_pre} (in a -backward manner) to allow a weaker precondition, then use the rules in @{attribute PrePostE_intro} -to derive a weakest precondition, and then use @{method auto} to show that it is implied by -the given precondition. For more serious proofs, one will want to set up specialised proof -tactics. This example uses only basic proof methods, to make the reasoning steps more explicit.\<close> +backward manner) to allow a weaker precondition, then use the rules in +@{attribute PrePostE_compositeI} and @{attribute PrePostE_atomI} to derive a weakest precondition, +and then use @{method auto} to show that it is implied by the given precondition. For more serious +proofs, one will want to set up specialised proof tactics. This example uses only basic proof +methods, to make the reasoning steps more explicit.\<close> (*<*) end diff --git a/snapshots/isabelle/aarch64/Aarch64.thy b/snapshots/isabelle/aarch64/Aarch64.thy index 55125942..a2156632 100644 --- a/snapshots/isabelle/aarch64/Aarch64.thy +++ b/snapshots/isabelle/aarch64/Aarch64.thy @@ -1,18 +1,18 @@ -chapter \<open>Generated by Lem from aarch64.lem.\<close> +chapter \<open>Generated by Lem from \<open>aarch64.lem\<close>.\<close> theory "Aarch64" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_operators_mwords" - "Sail2_prompt_monad" - "Sail2_prompt" - "Sail2_string" - "Aarch64_types" - "Aarch64_extras" +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" + "Aarch64_types" + "Aarch64_extras" begin @@ -27,6 +27,12 @@ begin (*open import Aarch64_types*) (*open import Aarch64_extras*) +(*val eq_unit : unit -> unit -> bool*) + +definition eq_unit :: " unit \<Rightarrow> unit \<Rightarrow> bool " where + " eq_unit g__306 g__307 = ( True )" + + @@ -141,7 +147,7 @@ definition ReadRAM :: "('m::len)itself \<Rightarrow> int \<Rightarrow>('m::len) definition WriteRAM :: "('m::len)itself \<Rightarrow> int \<Rightarrow>('m::len)Word.word \<Rightarrow>('m::len)Word.word \<Rightarrow>('p8_times_n_::len)Word.word \<Rightarrow>((register_value),(unit),(exception))monad " where " WriteRAM addr_length bytes hex_ram addr data = ( (let addr_length = (size_itself_int addr_length) in - write_ram addr_length bytes hex_ram addr data))" + write_ram addr_length bytes hex_ram addr data \<then> return () ))" definition TraceMemoryWrite :: " int \<Rightarrow>('m::len)Word.word \<Rightarrow>('p8_times_n_::len)Word.word \<Rightarrow> unit " where @@ -170,7 +176,7 @@ definition extsv :: " int \<Rightarrow>('n::len)Word.word \<Rightarrow>('m::len definition slice_mask :: " int \<Rightarrow> int \<Rightarrow> int \<Rightarrow>('n::len)Word.word " where " slice_mask (n__tv :: int) i l = ( - (let (one :: 'n bits) = ((extzv n__tv (vec_of_bits [B1] :: 1 Word.word) :: ( 'n::len)Word.word)) in + (let one = ((extzv n__tv (vec_of_bits [B1] :: 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)))" @@ -187,8 +193,7 @@ definition is_zero_subrange :: "('n::len)Word.word \<Rightarrow> int \<Rightarr definition is_ones_subrange :: "('n::len)Word.word \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool " where " is_ones_subrange xs i j = ( - (let (m :: 'n bits) = - ((slice_mask ((int (size xs))) j ((((j - i)) + (( 1 :: int)::ii))) :: ( 'n::len)Word.word)) in + (let m = ((slice_mask ((int (size xs))) j ((((j - i)) + (( 1 :: int)::ii))) :: ( 'n::len)Word.word)) in (((and_vec xs m :: ( 'n::len)Word.word)) = m)))" @@ -325,7 +330,7 @@ definition unsigned_subrange :: "('n::len)Word.word \<Rightarrow> int \<Rightar definition zext_ones :: " int \<Rightarrow> int \<Rightarrow>('n::len)Word.word " where " zext_ones (n__tv :: int) m = ( - (let (v :: 'n bits) = ((extsv n__tv (vec_of_bits [B1] :: 1 Word.word) :: ( 'n::len)Word.word)) in + (let v = ((extsv n__tv (vec_of_bits [B1] :: 1 Word.word) :: ( 'n::len)Word.word)) in (shiftr v ((((int (size v))) - m)) :: ( 'n::len)Word.word)))" @@ -733,28 +738,28 @@ definition undefined_InstrSet :: " unit \<Rightarrow>((register_value),(InstrSe definition undefined_ProcState :: " unit \<Rightarrow>((register_value),(ProcState),(exception))monad " where " undefined_ProcState _ = ( - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__0 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__2 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__3 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__4 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__5 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__6 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__7 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__8 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__9 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__10 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__11 :: 1 bits) . - (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__12 :: 2 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__13 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__14 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__15 :: 1 bits) . - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__16 :: 4 bits) . - (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (w__17 :: 8 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__18 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__19 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__20 :: 1 bits) . - (undefined_bitvector (( 5 :: int)::ii) :: ( 5 Word.word) M) \<bind> (\<lambda> (w__21 :: 5 bits) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__0 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__2 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__3 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__4 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__5 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__6 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__7 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__8 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__9 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__10 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__11 :: 1 Word.word) . + (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__12 :: 2 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__13 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__14 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__15 :: 1 Word.word) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__16 :: 4 Word.word) . + (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (w__17 :: 8 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__18 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__19 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__20 :: 1 Word.word) . + (undefined_bitvector (( 5 :: int)::ii) :: ( 5 Word.word) M) \<bind> (\<lambda> (w__21 :: 5 Word.word) . return ((| ProcState_N = w__0, ProcState_Z = w__1, ProcState_C = w__2, @@ -818,10 +823,10 @@ definition undefined_BranchType :: " unit \<Rightarrow>((register_value),(Branc definition undefined_ExceptionRecord :: " unit \<Rightarrow>((register_value),(ExceptionRecord),(exception))monad " where " undefined_ExceptionRecord _ = ( undefined_Exception () \<bind> (\<lambda> (w__0 :: Exception) . - (undefined_bitvector (( 25 :: int)::ii) :: ( 25 Word.word) M) \<bind> (\<lambda> (w__1 :: 25 bits) . - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . + (undefined_bitvector (( 25 :: int)::ii) :: ( 25 Word.word) M) \<bind> (\<lambda> (w__1 :: 25 Word.word) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . undefined_bool () \<bind> (\<lambda> (w__3 :: bool) . - (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__4 :: 52 bits) . + (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__4 :: 52 Word.word) . return ((| ExceptionRecord_typ = w__0, ExceptionRecord_syndrome = w__1, ExceptionRecord_vaddress = w__2, @@ -945,15 +950,15 @@ definition undefined_FaultRecord :: " unit \<Rightarrow>((register_value),(Faul " undefined_FaultRecord _ = ( undefined_Fault () \<bind> (\<lambda> (w__0 :: Fault) . undefined_AccType () \<bind> (\<lambda> (w__1 :: AccType) . - (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__2 :: 52 bits) . + (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__2 :: 52 Word.word) . undefined_bool () \<bind> (\<lambda> (w__3 :: bool) . undefined_bool () \<bind> (\<lambda> (w__4 :: bool) . undefined_int () \<bind> (\<lambda> (w__5 :: ii) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__6 :: 1 bits) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__6 :: 1 Word.word) . undefined_bool () \<bind> (\<lambda> (w__7 :: bool) . - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__8 :: 4 bits) . - (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__9 :: 2 bits) . - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__10 :: 4 bits) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__8 :: 4 Word.word) . + (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__9 :: 2 Word.word) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__10 :: 4 Word.word) . return ((| FaultRecord_typ = w__0, FaultRecord_acctype = w__1, FaultRecord_ipaddress = w__2, @@ -1072,8 +1077,8 @@ definition undefined_DeviceType :: " unit \<Rightarrow>((register_value),(Devic definition undefined_MemAttrHints :: " unit \<Rightarrow>((register_value),(MemAttrHints),(exception))monad " where " undefined_MemAttrHints _ = ( - (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__0 :: 2 bits) . - (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__1 :: 2 bits) . + (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__0 :: 2 Word.word) . + (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__1 :: 2 Word.word) . undefined_bool () \<bind> (\<lambda> (w__2 :: bool) . return ((| MemAttrHints_attrs = w__0, MemAttrHints_hints = w__1, @@ -1102,8 +1107,8 @@ definition undefined_MemoryAttributes :: " unit \<Rightarrow>((register_value), definition undefined_FullAddress :: " unit \<Rightarrow>((register_value),(FullAddress),(exception))monad " where " undefined_FullAddress _ = ( - (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__0 :: 52 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 bits) . + (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__0 :: 52 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 Word.word) . return ((| FullAddress_physicaladdress = w__0, FullAddress_NS = w__1 |)))))" @@ -1115,7 +1120,7 @@ definition undefined_AddressDescriptor :: " unit \<Rightarrow>((register_value) undefined_FaultRecord () \<bind> (\<lambda> (w__0 :: FaultRecord) . undefined_MemoryAttributes () \<bind> (\<lambda> (w__1 :: MemoryAttributes) . undefined_FullAddress () \<bind> (\<lambda> (w__2 :: FullAddress) . - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . return ((| AddressDescriptor_fault = w__0, AddressDescriptor_memattrs = w__1, AddressDescriptor_paddress = w__2, @@ -1519,10 +1524,10 @@ definition undefined_AccessDescriptor :: " unit \<Rightarrow>((register_value), definition undefined_Permissions :: " unit \<Rightarrow>((register_value),(Permissions),(exception))monad " where " undefined_Permissions _ = ( - (undefined_bitvector (( 3 :: int)::ii) :: ( 3 Word.word) M) \<bind> (\<lambda> (w__0 :: 3 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__2 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__3 :: 1 bits) . + (undefined_bitvector (( 3 :: int)::ii) :: ( 3 Word.word) M) \<bind> (\<lambda> (w__0 :: 3 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__2 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__3 :: 1 Word.word) . return ((| Permissions_ap = w__0, Permissions_xn = w__1, Permissions_xxn = w__2, @@ -1534,13 +1539,13 @@ definition undefined_Permissions :: " unit \<Rightarrow>((register_value),(Perm definition undefined_TLBRecord :: " unit \<Rightarrow>((register_value),(TLBRecord),(exception))monad " where " undefined_TLBRecord _ = ( undefined_Permissions () \<bind> (\<lambda> (w__0 :: Permissions) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 bits) . - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__2 :: 4 bits) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 Word.word) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__2 :: 4 Word.word) . undefined_bool () \<bind> (\<lambda> (w__3 :: bool) . undefined_int () \<bind> (\<lambda> (w__4 :: ii) . undefined_int () \<bind> (\<lambda> (w__5 :: ii) . undefined_DescriptorUpdate () \<bind> (\<lambda> (w__6 :: DescriptorUpdate) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__7 :: 1 bits) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__7 :: 1 Word.word) . undefined_AddressDescriptor () \<bind> (\<lambda> (w__8 :: AddressDescriptor) . return ((| TLBRecord_perms = w__0, TLBRecord_nG = w__1, @@ -1709,8 +1714,8 @@ definition undefined_PrivilegeLevel :: " unit \<Rightarrow>((register_value),(P definition undefined_AArch32_SErrorSyndrome :: " unit \<Rightarrow>((register_value),(AArch32_SErrorSyndrome),(exception))monad " where " undefined_AArch32_SErrorSyndrome _ = ( - (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__0 :: 2 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 bits) . + (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__0 :: 2 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 Word.word) . return ((| AArch32_SErrorSyndrome_AET = w__0, AArch32_SErrorSyndrome_ExT = w__1 |)))))" @@ -1748,13 +1753,13 @@ definition undefined_SystemOp :: " unit \<Rightarrow>((register_value),(SystemO definition undefined_PCSample :: " unit \<Rightarrow>((register_value),(PCSample),(exception))monad " where " undefined_PCSample _ = ( undefined_bool () \<bind> (\<lambda> (w__0 :: bool) . - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . - (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__2 :: 2 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__3 :: 1 bits) . - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__4 :: 1 bits) . - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 bits) . - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . - (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__2 :: 2 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__3 :: 1 Word.word) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__4 :: 1 Word.word) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 Word.word) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . + (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 Word.word) . return ((| PCSample_valid_name = w__0, PCSample_pc = w__1, PCSample_el = w__2, @@ -3288,9 +3293,9 @@ definition AArch64_CreateFaultRecord :: " Fault \<Rightarrow>(52)Word.word \<Ri " AArch64_CreateFaultRecord typ1 ipaddress level acctype write1 extflag errortype secondstage s2fs1walk = ( undefined_FaultRecord () \<bind> (\<lambda> (fault :: FaultRecord) . (let fault = ((fault (| FaultRecord_typ := typ1 |))) in - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__0 :: 4 bits) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__0 :: 4 Word.word) . (let fault = ((fault (| FaultRecord_domain := w__0 |))) in - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__1 :: 4 bits) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__1 :: 4 Word.word) . (let (fault :: FaultRecord) = ((fault (| FaultRecord_debugmoe := w__1 |))) in (let (fault :: FaultRecord) = ((fault (| FaultRecord_errortype := errortype |))) in (let (fault :: FaultRecord) = ((fault (| FaultRecord_ipaddress := ipaddress |))) in @@ -3747,7 +3752,7 @@ definition CombineS1S2AttrHints :: " MemAttrHints \<Rightarrow> MemAttrHints \< " CombineS1S2AttrHints s1desc s2desc = ( undefined_MemAttrHints () \<bind> (\<lambda> (result :: MemAttrHints) . (if (((((((MemAttrHints_attrs s2desc) = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<or> ((((MemAttrHints_attrs s1desc) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then - (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__0 :: 2 bits) . + (undefined_bitvector (( 2 :: int)::ii) :: ( 2 Word.word) M) \<bind> (\<lambda> (w__0 :: 2 Word.word) . (let (result :: MemAttrHints) = ((result (| MemAttrHints_attrs := w__0 |))) in return result)) else @@ -3814,7 +3819,7 @@ definition aget_Vpart :: " int \<Rightarrow> int \<Rightarrow> int \<Rightarrow (assert_exp (((width__tv = (( 64 :: int)::ii)))) (''(width == 64)'') \<then> read_reg V_ref) \<bind> (\<lambda> (w__1 :: ( 128 bits) list) . return ((Word.ucast - ((slice ((access_list_dec w__1 n :: 128 Word.word)) (( 64 :: int)::ii) (( 64 :: int)::ii) :: ( 'width::len)Word.word)) + ((slice ((access_list_dec w__1 n :: 128 Word.word)) (( 64 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) :: ( 'width::len)Word.word)))))" @@ -3885,8 +3890,8 @@ definition CountLeadingSignBits :: "('N::len)Word.word \<Rightarrow>((register_ definition BitReverse :: "('N::len)Word.word \<Rightarrow>((register_value),(('N::len)Word.word),(exception))monad " where " BitReverse data = ( - (undefined_bitvector ((int (size data))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . - (let (result :: 'N bits) = + (undefined_bitvector ((int (size data))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . + (let result = (foreach (index_list (( 0 :: int)::ii) ((((int (size data))) - (( 1 :: int)::ii))) (( 1 :: int)::ii)) result (\<lambda> i result . (set_slice ((int (size data))) (( 1 :: int)::ii) result @@ -4012,7 +4017,7 @@ definition FPNeg :: "('N::len)Word.word \<Rightarrow>((register_value),(('N::le :: 1 Word.word) :: 1 Word.word)) ((slice op1 (( 0 :: int)::ii) (((( 16 :: int)::ii) - (( 1 :: int)::ii))) :: 15 Word.word)) - :: ( 'N::len)Word.word)) + :: 16 Word.word)) :: ( 'N::len)Word.word))) else if (((p00 = (( 32 :: int)::ii)))) then (let (op1 :: 32 Word.word) = ((Word.ucast op1 :: 32 Word.word)) in @@ -4024,7 +4029,7 @@ definition FPNeg :: "('N::len)Word.word \<Rightarrow>((register_value),(('N::le :: 1 Word.word) :: 1 Word.word)) ((slice op1 (( 0 :: int)::ii) (((( 32 :: int)::ii) - (( 1 :: int)::ii))) :: 31 Word.word)) - :: ( 'N::len)Word.word)) + :: 32 Word.word)) :: ( 'N::len)Word.word))) else if (((p00 = (( 64 :: int)::ii)))) then (let (op1 :: 64 Word.word) = ((Word.ucast op1 :: 64 Word.word)) in @@ -4036,7 +4041,7 @@ definition FPNeg :: "('N::len)Word.word \<Rightarrow>((register_value),(('N::le :: 1 Word.word) :: 1 Word.word)) ((slice op1 (( 0 :: int)::ii) (((( 64 :: int)::ii) - (( 1 :: int)::ii))) :: 63 Word.word)) - :: ( 'N::len)Word.word)) + :: 64 Word.word)) :: ( 'N::len)Word.word))) else assert_exp False (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> exit0 () ))" @@ -4052,7 +4057,7 @@ definition FPAbs :: "('N::len)Word.word \<Rightarrow>((register_value),(('N::le return ((Word.ucast ((concat_vec (vec_of_bits [B0] :: 1 Word.word) ((slice op1 (( 0 :: int)::ii) (((( 16 :: int)::ii) - (( 1 :: int)::ii))) :: 15 Word.word)) - :: ( 'N::len)Word.word)) + :: 16 Word.word)) :: ( 'N::len)Word.word))) else if (((p00 = (( 32 :: int)::ii)))) then (let (op1 :: 32 Word.word) = ((Word.ucast op1 :: 32 Word.word)) in @@ -4060,7 +4065,7 @@ definition FPAbs :: "('N::len)Word.word \<Rightarrow>((register_value),(('N::le return ((Word.ucast ((concat_vec (vec_of_bits [B0] :: 1 Word.word) ((slice op1 (( 0 :: int)::ii) (((( 32 :: int)::ii) - (( 1 :: int)::ii))) :: 31 Word.word)) - :: ( 'N::len)Word.word)) + :: 32 Word.word)) :: ( 'N::len)Word.word))) else if (((p00 = (( 64 :: int)::ii)))) then (let (op1 :: 64 Word.word) = ((Word.ucast op1 :: 64 Word.word)) in @@ -4068,7 +4073,7 @@ definition FPAbs :: "('N::len)Word.word \<Rightarrow>((register_value),(('N::le return ((Word.ucast ((concat_vec (vec_of_bits [B0] :: 1 Word.word) ((slice op1 (( 0 :: int)::ii) (((( 64 :: int)::ii) - (( 1 :: int)::ii))) :: 63 Word.word)) - :: ( 'N::len)Word.word)) + :: 64 Word.word)) :: ( 'N::len)Word.word))) else assert_exp False (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> exit0 () ))" @@ -4178,22 +4183,22 @@ definition AArch32_ReportHypEntry :: " ExceptionRecord \<Rightarrow>((register_ :: 32 Word.word)) \<then> (if ((((((typ1 = Exception_InstructionAbort))) \<or> (((typ1 = Exception_PCAlignment)))))) then (write_reg HIFAR_ref ((slice(ExceptionRecord_vaddress exception) (( 0 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__0 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__0 :: 32 Word.word) . write_reg HDFAR_ref w__0) else if (((typ1 = Exception_DataAbort))) then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . write_reg HIFAR_ref w__1 \<then> write_reg HDFAR_ref ((slice(ExceptionRecord_vaddress exception) (( 0 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word))) else return () )) \<then> (if(ExceptionRecord_ipavalid exception) then - (read_reg HPFAR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg HPFAR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . write_reg HPFAR_ref ((set_slice (( 32 :: int)::ii) (( 28 :: int)::ii) w__2 (( 4 :: int)::ii) ((slice(ExceptionRecord_ipaddress exception) (( 12 :: int)::ii) (( 28 :: int)::ii) :: 28 Word.word)) :: 32 Word.word))) else - (read_reg HPFAR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__3 :: 32 Word.word) . + (read_reg HPFAR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__3 :: 32 bits) . (undefined_bitvector (( 28 :: int)::ii) :: ( 28 Word.word) M) \<bind> (\<lambda> (w__4 :: 28 Word.word) . write_reg HPFAR_ref ((set_slice (( 32 :: int)::ii) (( 28 :: int)::ii) w__3 (( 4 :: int)::ii) w__4 :: 32 Word.word)))))))))))))))" @@ -4216,8 +4221,7 @@ definition aset_Elem__0 :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>( definition aset_Elem__1 :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>('size::len)Word.word \<Rightarrow>((register_value),(('N::len)Word.word),(exception))monad " where " aset_Elem__1 vector_name__arg e value_name = ( (let vector_name = vector_name__arg in - (aset_Elem__0 vector_name e ((make_the_value ((int (size value_name))) :: ( 'size::len)itself)) value_name - :: (( 'N::len)Word.word) M)))" + (aset_Elem__0 vector_name e ((make_the_value ((int (size value_name))) )) value_name :: (( 'N::len)Word.word) M)))" (*val aget_Elem__0 : forall 'N 'size . Size 'N, Size 'size => mword 'N -> ii -> itself 'size -> M (mword 'size)*) @@ -4233,7 +4237,7 @@ definition aget_Elem__0 :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>( definition aget_Elem__1 :: " int \<Rightarrow>('N::len)Word.word \<Rightarrow> int \<Rightarrow>((register_value),(('size::len)Word.word),(exception))monad " where " aget_Elem__1 (size__tv :: int) vector_name e = ( - (aget_Elem__0 vector_name e ((make_the_value size__tv :: ( 'size::len)itself)) :: (( 'size::len)Word.word) M))" + (aget_Elem__0 vector_name e ((make_the_value size__tv )) :: (( 'size::len)Word.word) M))" (*val UnsignedSatQ : forall 'N . Size 'N => ii -> itself 'N -> M (mword 'N * bool)*) @@ -4259,7 +4263,7 @@ definition UnsignedSatQ :: " int \<Rightarrow>('N::len)itself \<Rightarrow>((re (let (saturated :: bool) = False in (result, saturated)))) in (result, saturated))) in - return ((GetSlice_int ((make_the_value N :: ( 'N::len)itself)) result (( 0 :: int)::ii) :: ( 'N::len)Word.word), saturated))))))" + return ((GetSlice_int ((make_the_value N )) result (( 0 :: int)::ii) :: ( 'N::len)Word.word), saturated))))))" (*val SignedSatQ : forall 'N . Size 'N => ii -> itself 'N -> M (mword 'N * bool)*) @@ -4285,7 +4289,7 @@ definition SignedSatQ :: " int \<Rightarrow>('N::len)itself \<Rightarrow>((regi (let (saturated :: bool) = False in (result, saturated)))) in (result, saturated))) in - return ((GetSlice_int ((make_the_value N :: ( 'N::len)itself)) result (( 0 :: int)::ii) :: ( 'N::len)Word.word), saturated))))))" + return ((GetSlice_int ((make_the_value N )) result (( 0 :: int)::ii) :: ( 'N::len)Word.word), saturated))))))" (*val SatQ : forall 'N . Size 'N => ii -> itself 'N -> bool -> M (mword 'N * bool)*) @@ -4294,10 +4298,10 @@ definition SatQ :: " int \<Rightarrow>('N::len)itself \<Rightarrow> bool \<Righ " SatQ i N unsigned = ( (let N = (size_itself_int N) in undefined_bool () \<bind> (\<lambda> (sat :: bool) . - (undefined_bitvector N :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . - (if unsigned then (UnsignedSatQ i ((make_the_value N :: ( 'N::len)itself)) :: ((( 'N::len)Word.word * bool)) M) - else (SignedSatQ i ((make_the_value N :: ( 'N::len)itself)) :: ((( 'N::len)Word.word * bool)) M)) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in - (let (result :: 'N bits) = tup__0 in + (undefined_bitvector N :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . + (if unsigned then (UnsignedSatQ i ((make_the_value N )) ) + else (SignedSatQ i ((make_the_value N )) )) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (let result = tup__0 in (let (sat :: bool) = tup__1 in return (result, sat)))))))))" @@ -4308,7 +4312,7 @@ definition Replicate :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow>((re " Replicate (N__tv :: int) x = ( assert_exp (((((N__tv mod ((int (size x))))) = (( 0 :: int)::ii)))) (''((N MOD M) == 0)'') \<then> ((let O1 = (N__tv div ((int (size x)))) in - assert_exp True ('''') \<then> return ((replicate_bits x ((N__tv div ((int (size x))))) :: ( 'N::len)Word.word)))))" + assert_exp True ('''') \<then> return ((replicate_bits x ((N__tv div ((int (size x))))))))))" (*val Zeros__0 : forall 'N . Size 'N => itself 'N -> mword 'N*) @@ -4322,14 +4326,14 @@ definition Zeros__0 :: "('N::len)itself \<Rightarrow>('N::len)Word.word " wher definition Zeros__1 :: " int \<Rightarrow> unit \<Rightarrow>('N::len)Word.word " where - " Zeros__1 (N__tv :: int) _ = ( (Zeros__0 ((make_the_value N__tv :: ( 'N::len)itself)) :: ( 'N::len)Word.word))" + " Zeros__1 (N__tv :: int) _ = ( (Zeros__0 ((make_the_value N__tv )) :: ( 'N::len)Word.word))" (*val __ResetMemoryState : unit -> M unit*) definition ResetMemoryState :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " ResetMemoryState _ = ( - (read_reg Memory_ref :: ( 52 Word.word) M) \<bind> (\<lambda> (w__0 :: 52 Word.word) . + (read_reg Memory_ref :: ( 52 Word.word) M) \<bind> (\<lambda> (w__0 :: 52 bits) . (let (_ :: unit) = (InitRAM (( 52 :: int)::ii) (( 1 :: int)::ii) w__0 ((Zeros__0 ((make_the_value (( 8 :: int)::ii) :: 8 itself)) :: 8 Word.word))) in write_reg ExclusiveLocal_ref False)))" @@ -4346,8 +4350,7 @@ definition ZeroExtend__0 :: "('M::len)Word.word \<Rightarrow>('N::len)itself \< definition ZeroExtend__1 :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow>((register_value),(('N::len)Word.word),(exception))monad " where - " ZeroExtend__1 (N__tv :: int) x = ( - (ZeroExtend__0 x ((make_the_value N__tv :: ( 'N::len)itself)) :: (( 'N::len)Word.word) M))" + " ZeroExtend__1 (N__tv :: int) x = ( (ZeroExtend__0 x ((make_the_value N__tv )) :: (( 'N::len)Word.word) M))" (*val aset_Vpart : forall 'width . Size 'width => ii -> ii -> mword 'width -> M unit*) @@ -4358,8 +4361,8 @@ definition aset_Vpart :: " int \<Rightarrow> int \<Rightarrow>('width::len)Word assert_exp ((((((part = (( 0 :: int)::ii)))) \<or> (((part = (( 1 :: int)::ii))))))) (''((part == 0) || (part == 1))'')) \<then> (if (((part = (( 0 :: int)::ii)))) then (assert_exp ((((((((int (size value_name))) = (( 8 :: int)::ii)))) \<or> ((((((((int (size value_name))) = (( 16 :: int)::ii)))) \<or> ((((((((int (size value_name))) = (( 32 :: int)::ii)))) \<or> (((((int (size value_name))) = (( 64 :: int)::ii))))))))))))) (''((width == 8) || ((width == 16) || ((width == 32) || (width == 64))))'') \<then> - read_reg V_ref) \<bind> (\<lambda> (w__0 :: ( 128 Word.word) list) . - (ZeroExtend__1 (( 128 :: int)::ii) value_name :: ( 128 Word.word) M) \<bind> (\<lambda> (w__1 :: 128 bits) . + read_reg V_ref) \<bind> (\<lambda> (w__0 :: ( 128 bits) list) . + (ZeroExtend__1 (( 128 :: int)::ii) value_name :: ( 128 Word.word) M) \<bind> (\<lambda> (w__1 :: 128 Word.word) . write_reg V_ref ((update_list_dec w__0 n w__1 :: ( 128 Word.word) list)))) else (assert_exp (((((int (size value_name))) = (( 64 :: int)::ii)))) (''(width == 64)'') \<then> @@ -4369,7 +4372,7 @@ definition aset_Vpart :: " int \<Rightarrow> int \<Rightarrow>('width::len)Word ((update_subrange_vec_dec tmp_2870 (( 127 :: int)::ii) (( 64 :: int)::ii) ((subrange_vec_dec value_name (( 63 :: int)::ii) (( 0 :: int)::ii) :: 64 Word.word)) :: 128 Word.word)) in - read_reg V_ref \<bind> (\<lambda> (w__3 :: ( 128 Word.word) list) . + read_reg V_ref \<bind> (\<lambda> (w__3 :: ( 128 bits) list) . write_reg V_ref ((update_list_dec w__3 n tmp_2870 :: ( 128 Word.word) list))))))))" @@ -4379,8 +4382,8 @@ definition aset_V :: " int \<Rightarrow>('width::len)Word.word \<Rightarrow>((r " aset_V n value_name = ( ((assert_exp (((((n \<ge> (( 0 :: int)::ii))) \<and> ((n \<le> (( 31 :: int)::ii)))))) (''((n >= 0) && (n <= 31))'') \<then> assert_exp ((((((((int (size value_name))) = (( 8 :: int)::ii)))) \<or> ((((((((int (size value_name))) = (( 16 :: int)::ii)))) \<or> ((((((((int (size value_name))) = (( 32 :: int)::ii)))) \<or> ((((((((int (size value_name))) = (( 64 :: int)::ii)))) \<or> (((((int (size value_name))) = (( 128 :: int)::ii)))))))))))))))) (''((width == 8) || ((width == 16) || ((width == 32) || ((width == 64) || (width == 128)))))'')) \<then> - read_reg V_ref) \<bind> (\<lambda> (w__0 :: ( 128 Word.word) list) . - (ZeroExtend__1 (( 128 :: int)::ii) value_name :: ( 128 Word.word) M) \<bind> (\<lambda> (w__1 :: 128 bits) . + read_reg V_ref) \<bind> (\<lambda> (w__0 :: ( 128 bits) list) . + (ZeroExtend__1 (( 128 :: int)::ii) value_name :: ( 128 Word.word) M) \<bind> (\<lambda> (w__1 :: 128 Word.word) . write_reg V_ref ((update_list_dec w__0 n w__1 :: ( 128 Word.word) list)))))" @@ -4390,7 +4393,7 @@ definition AArch64_ResetSIMDFPRegisters :: " unit \<Rightarrow>((register_value " AArch64_ResetSIMDFPRegisters _ = ( (foreachM (index_list (( 0 :: int)::ii) (( 31 :: int)::ii) (( 1 :: int)::ii)) () (\<lambda> i unit_var . - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . aset_V i w__0))))" + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . aset_V i w__0))))" (*val aset_SP : forall 'width . Size 'width => mword 'width -> M unit*) @@ -4400,23 +4403,23 @@ definition aset_SP :: "('width::len)Word.word \<Rightarrow>((register_value),(u (assert_exp ((((((((int (size value_name))) = (( 32 :: int)::ii)))) \<or> (((((int (size value_name))) = (( 64 :: int)::ii))))))) (''((width == 32) || (width == 64))'') \<then> read_reg PSTATE_ref) \<bind> (\<lambda> (w__0 :: ProcState) . if ((((ProcState_SP w__0) = (vec_of_bits [B0] :: 1 Word.word)))) then - (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . + (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . write_reg SP_EL0_ref w__1) else read_reg PSTATE_ref \<bind> (\<lambda> (w__2 :: ProcState) . (let p__299 = ((ProcState_EL w__2)) in (let pat0 = p__299 in if (((pat0 = EL0))) then - (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 bits) . + (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . write_reg SP_EL0_ref w__3) else if (((pat0 = EL1))) then - (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 bits) . + (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 Word.word) . write_reg SP_EL1_ref w__4) else if (((pat0 = EL2))) then - (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__5 :: 64 bits) . + (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__5 :: 64 Word.word) . write_reg SP_EL2_ref w__5) else - (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: 64 bits) . + (ZeroExtend__1 (( 64 :: int)::ii) value_name :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: 64 Word.word) . write_reg SP_EL3_ref w__6))))))" @@ -4425,7 +4428,7 @@ definition aset_SP :: "('width::len)Word.word \<Rightarrow>((register_value),(u definition LSR_C :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>((register_value),(('N::len)Word.word*(1)Word.word),(exception))monad " where " LSR_C x shift = ( assert_exp ((shift > (( 0 :: int)::ii))) (''(shift > 0)'') \<then> - ((let (result :: 'N bits) = ((shiftr x shift :: ( 'N::len)Word.word)) in + ((let result = ((shiftr x shift :: ( 'N::len)Word.word)) in (let (carry_out :: 1 bits) = (if ((shift > ((int (size result))))) then (vec_of_bits [B0] :: 1 Word.word) else (vec_of_bits [access_vec_dec x ((shift - (( 1 :: int)::ii)))] :: 1 Word.word)) in @@ -4438,11 +4441,11 @@ definition LSR :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>((register " LSR x shift = ( (assert_exp ((shift \<ge> (( 0 :: int)::ii))) (''(shift >= 0)'') \<then> (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (anon10 :: 1 bits) . - (undefined_bitvector ((int (size x))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size x))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . if (((shift = (( 0 :: int)::ii)))) then return x else - (LSR_C x shift :: ((( 'N::len)Word.word * 1 Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in - (let (result :: 'N bits) = tup__0 in + (LSR_C x shift ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (let result = tup__0 in (let (anon10 :: 1 bits) = tup__1 in return result)))))))" @@ -4453,7 +4456,7 @@ definition Poly32Mod2 :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Right " Poly32Mod2 data__arg poly = ( (let data = data__arg in assert_exp ((((int (size data))) > (( 32 :: int)::ii))) (''(N > 32)'') \<then> - ((let (poly' :: 'N bits) = ((extzv ((int (size data))) poly :: ( 'N::len)Word.word)) in + ((let poly' = ((extzv ((int (size data))) poly :: ( 'N::len)Word.word)) in (let (data :: ( 'N::len)Word.word) = (foreach (index_list ((((int (size data))) - (( 1 :: int)::ii))) (( 32 :: int)::ii) (- (( 1 :: int)::ii))) data (\<lambda> i data . @@ -4469,7 +4472,7 @@ definition Poly32Mod2 :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Right definition LSL_C :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>((register_value),(('N::len)Word.word*(1)Word.word),(exception))monad " where " LSL_C x shift = ( assert_exp ((shift > (( 0 :: int)::ii))) (''(shift > 0)'') \<then> - ((let (result :: 'N bits) = ((shiftl x shift :: ( 'N::len)Word.word)) in + ((let result = ((shiftl x shift :: ( 'N::len)Word.word)) in (let (carry_out :: 1 bits) = (if ((shift > ((int (size result))))) then (vec_of_bits [B0] :: 1 Word.word) else (vec_of_bits [access_vec_dec x ((((int (size result))) - shift))] :: 1 Word.word)) in @@ -4482,11 +4485,11 @@ definition LSL :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>((register " LSL x shift = ( (assert_exp ((shift \<ge> (( 0 :: int)::ii))) (''(shift >= 0)'') \<then> (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (anon10 :: 1 bits) . - (undefined_bitvector ((int (size x))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size x))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . if (((shift = (( 0 :: int)::ii)))) then return x else - (LSL_C x shift :: ((( 'N::len)Word.word * 1 Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in - (let (result :: 'N bits) = tup__0 in + (LSL_C x shift ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (let result = tup__0 in (let (anon10 :: 1 bits) = tup__1 in return result)))))))" @@ -4523,7 +4526,7 @@ definition LSInstructionSyndrome :: " unit \<Rightarrow>((register_value),((11) (*val IsZero : forall 'N . Size 'N => mword 'N -> bool*) definition IsZero :: "('N::len)Word.word \<Rightarrow> bool " where - " IsZero x = ( (x = ((Zeros__0 ((make_the_value ((int (size x))) :: ( 'N::len)itself)) :: ( 'N::len)Word.word))))" + " IsZero x = ( (x = ((Zeros__0 ((make_the_value ((int (size x))) )) :: ( 'N::len)Word.word))))" (*val IsZeroBit : forall 'N . Size 'N => mword 'N -> mword ty1*) @@ -4540,8 +4543,7 @@ definition AddWithCarry :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word " AddWithCarry x y carry_in = ( (let (unsigned_sum :: ii) = (((((Word.uint x)) + ((Word.uint y)))) + ((Word.uint carry_in))) in (let (signed_sum :: ii) = (((((Word.sint x)) + ((Word.sint y)))) + ((Word.uint carry_in))) in - (let (result :: 'N bits) = - ((GetSlice_int ((make_the_value ((int (size x))) :: ( 'N::len)itself)) unsigned_sum (( 0 :: int)::ii) :: ( 'N::len)Word.word)) in + (let result = ((GetSlice_int ((make_the_value ((int (size x))) )) unsigned_sum (( 0 :: int)::ii) :: ( 'N::len)Word.word)) in (let (n :: 1 bits) = ((vec_of_bits [access_vec_dec result ((((int (size result))) - (( 1 :: int)::ii)))] :: 1 Word.word)) in (let (z :: 1 bits) = @@ -4634,23 +4636,23 @@ definition FPZero :: " int \<Rightarrow>(1)Word.word \<Rightarrow>((register_va if (((p00 = (( 16 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 16 :: int)::ii) - (( 5 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 5 bits) = ((Zeros__0 ((make_the_value (( 5 :: int)::ii) :: 5 itself)) :: 5 Word.word)) in + (let (exp1 :: 5 bits) = ((Zeros__0 ((make_the_value (( 5 :: int)::ii) :: 5 itself)) :: 5 Word.word)) in (let (frac :: 10 bits) = ((Zeros__0 ((make_the_value F :: 10 itself)) :: 10 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 6 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 6 Word.word)) frac :: 16 Word.word)) :: ( 'N::len)Word.word)))))) else if (((p00 = (( 32 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 32 :: int)::ii) - (( 8 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 8 bits) = ((Zeros__0 ((make_the_value (( 8 :: int)::ii) :: 8 itself)) :: 8 Word.word)) in + (let (exp1 :: 8 bits) = ((Zeros__0 ((make_the_value (( 8 :: int)::ii) :: 8 itself)) :: 8 Word.word)) in (let (frac :: 23 bits) = ((Zeros__0 ((make_the_value F :: 23 itself)) :: 23 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 9 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 9 Word.word)) frac :: 32 Word.word)) :: ( 'N::len)Word.word)))))) else if (((p00 = (( 64 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 64 :: int)::ii) - (( 11 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 11 bits) = ((Zeros__0 ((make_the_value (( 11 :: int)::ii) :: 11 itself)) :: 11 Word.word)) in + (let (exp1 :: 11 bits) = ((Zeros__0 ((make_the_value (( 11 :: int)::ii) :: 11 itself)) :: 11 Word.word)) in (let (frac :: 52 bits) = ((Zeros__0 ((make_the_value F :: 52 itself)) :: 52 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 12 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 12 Word.word)) frac :: 64 Word.word)) :: ( 'N::len)Word.word)))))) else assert_exp False (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> exit0 () ))" @@ -4677,7 +4679,7 @@ definition ConstrainUnpredictableBits :: " int \<Rightarrow> Unpredictable \<Ri " ConstrainUnpredictableBits (width__tv :: int) which = ( (let (c :: Constraint) = (ConstrainUnpredictable which) in if (((c = Constraint_UNKNOWN))) then - return (c, (Zeros__0 ((make_the_value width__tv :: ( 'width::len)itself)) :: ( 'width::len)Word.word)) + return (c, (Zeros__0 ((make_the_value width__tv )) :: ( 'width::len)Word.word)) else (undefined_bitvector width__tv :: (( 'width::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'width::len)Word.word) . return (c, w__0))))" @@ -4725,7 +4727,7 @@ definition VFPExpandImm :: " int \<Rightarrow>(8)Word.word \<Rightarrow>((regis assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 16 :: int)::ii) - (( 5 :: int)::ii))) - (( 1 :: int)::ii)) in (let (sign :: 1 bits) = ((vec_of_bits [access_vec_dec imm8 (( 7 :: int)::ii)] :: 1 Word.word)) in - (let (exp :: 5 bits) = + (let (exp1 :: 5 bits) = ((concat_vec ((concat_vec ((not_vec (vec_of_bits [access_vec_dec imm8 (( 6 :: int)::ii)] :: 1 Word.word) :: 1 Word.word)) @@ -4738,13 +4740,13 @@ definition VFPExpandImm :: " int \<Rightarrow>(8)Word.word \<Rightarrow>((regis ((concat_vec ((subrange_vec_dec imm8 (( 3 :: int)::ii) (( 0 :: int)::ii) :: 4 Word.word)) ((Zeros__0 ((make_the_value ((F - (( 4 :: int)::ii))) :: 6 itself)) :: 6 Word.word)) :: 10 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 6 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 6 Word.word)) frac :: 16 Word.word)) :: ( 'N::len)Word.word))))))) else if (((p00 = (( 32 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 32 :: int)::ii) - (( 8 :: int)::ii))) - (( 1 :: int)::ii)) in (let (sign :: 1 bits) = ((vec_of_bits [access_vec_dec imm8 (( 7 :: int)::ii)] :: 1 Word.word)) in - (let (exp :: 8 bits) = + (let (exp1 :: 8 bits) = ((concat_vec ((concat_vec ((not_vec (vec_of_bits [access_vec_dec imm8 (( 6 :: int)::ii)] :: 1 Word.word) :: 1 Word.word)) @@ -4757,13 +4759,13 @@ definition VFPExpandImm :: " int \<Rightarrow>(8)Word.word \<Rightarrow>((regis ((concat_vec ((subrange_vec_dec imm8 (( 3 :: int)::ii) (( 0 :: int)::ii) :: 4 Word.word)) ((Zeros__0 ((make_the_value ((F - (( 4 :: int)::ii))) :: 19 itself)) :: 19 Word.word)) :: 23 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 9 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 9 Word.word)) frac :: 32 Word.word)) :: ( 'N::len)Word.word))))))) else if (((p00 = (( 64 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 64 :: int)::ii) - (( 11 :: int)::ii))) - (( 1 :: int)::ii)) in (let (sign :: 1 bits) = ((vec_of_bits [access_vec_dec imm8 (( 7 :: int)::ii)] :: 1 Word.word)) in - (let (exp :: 11 bits) = + (let (exp1 :: 11 bits) = ((concat_vec ((concat_vec ((not_vec (vec_of_bits [access_vec_dec imm8 (( 6 :: int)::ii)] :: 1 Word.word) :: 1 Word.word)) @@ -4776,7 +4778,7 @@ definition VFPExpandImm :: " int \<Rightarrow>(8)Word.word \<Rightarrow>((regis ((concat_vec ((subrange_vec_dec imm8 (( 3 :: int)::ii) (( 0 :: int)::ii) :: 4 Word.word)) ((Zeros__0 ((make_the_value ((F - (( 4 :: int)::ii))) :: 48 itself)) :: 48 Word.word)) :: 52 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 12 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 12 Word.word)) frac :: 64 Word.word)) :: ( 'N::len)Word.word))))))) else assert_exp False (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> exit0 () ))" @@ -4792,8 +4794,7 @@ definition SignExtend__0 :: "('M::len)Word.word \<Rightarrow>('N::len)itself \< definition SignExtend__1 :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow>((register_value),(('N::len)Word.word),(exception))monad " where - " SignExtend__1 (N__tv :: int) x = ( - (SignExtend__0 x ((make_the_value N__tv :: ( 'N::len)itself)) :: (( 'N::len)Word.word) M))" + " SignExtend__1 (N__tv :: int) x = ( (SignExtend__0 x ((make_the_value N__tv )) :: (( 'N::len)Word.word) M))" (*val Extend__0 : forall 'M 'N . Size 'M, Size 'N => mword 'M -> itself 'N -> bool -> M (mword 'N)*) @@ -4803,13 +4804,13 @@ definition SignExtend__1 :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow> definition Extend__0 :: "('M::len)Word.word \<Rightarrow>('N::len)itself \<Rightarrow> bool \<Rightarrow>((register_value),(('N::len)Word.word),(exception))monad " where " Extend__0 x N unsigned = ( (let N = (size_itself_int N) in - if unsigned then (ZeroExtend__0 x ((make_the_value N :: ( 'N::len)itself)) :: (( 'N::len)Word.word) M) - else (SignExtend__0 x ((make_the_value N :: ( 'N::len)itself)) :: (( 'N::len)Word.word) M)))" + if unsigned then (ZeroExtend__0 x ((make_the_value N )) :: (( 'N::len)Word.word) M) + else (SignExtend__0 x ((make_the_value N )) :: (( 'N::len)Word.word) M)))" definition Extend__1 :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow> bool \<Rightarrow>((register_value),(('N::len)Word.word),(exception))monad " where " Extend__1 (N__tv :: int) x unsigned = ( - (Extend__0 x ((make_the_value N__tv :: ( 'N::len)itself)) unsigned :: (( 'N::len)Word.word) M))" + (Extend__0 x ((make_the_value N__tv )) unsigned :: (( 'N::len)Word.word) M))" (*val ASR_C : forall 'N . Size 'N => mword 'N -> ii -> M (mword 'N * mword ty1)*) @@ -4817,7 +4818,7 @@ definition Extend__1 :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow> boo definition ASR_C :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>((register_value),(('N::len)Word.word*(1)Word.word),(exception))monad " where " ASR_C x shift = ( assert_exp ((shift > (( 0 :: int)::ii))) (''(shift > 0)'') \<then> - ((let (result :: 'N bits) = ((arith_shiftr x shift :: ( 'N::len)Word.word)) in + ((let result = ((arith_shiftr x shift :: ( 'N::len)Word.word)) in (let (carry_out :: 1 bits) = (if ((shift > ((int (size result))))) then (vec_of_bits [access_vec_dec x ((((int (size result))) - (( 1 :: int)::ii)))] :: 1 Word.word) @@ -4831,11 +4832,11 @@ definition ASR :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>((register " ASR x shift = ( (assert_exp ((shift \<ge> (( 0 :: int)::ii))) (''(shift >= 0)'') \<then> (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (anon10 :: 1 bits) . - (undefined_bitvector ((int (size x))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size x))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . if (((shift = (( 0 :: int)::ii)))) then return x else - (ASR_C x shift :: ((( 'N::len)Word.word * 1 Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in - (let (result :: 'N bits) = tup__0 in + (ASR_C x shift ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (let result = tup__0 in (let (anon10 :: 1 bits) = tup__1 in return result)))))))" @@ -4851,13 +4852,13 @@ definition Ones__0 :: "('N::len)itself \<Rightarrow>('N::len)Word.word " where definition Ones__1 :: " int \<Rightarrow> unit \<Rightarrow>('N::len)Word.word " where - " Ones__1 (N__tv :: int) _ = ( (Ones__0 ((make_the_value N__tv :: ( 'N::len)itself)) :: ( 'N::len)Word.word))" + " Ones__1 (N__tv :: int) _ = ( (Ones__0 ((make_the_value N__tv )) :: ( 'N::len)Word.word))" (*val IsOnes : forall 'N . Size 'N => mword 'N -> bool*) definition IsOnes :: "('N::len)Word.word \<Rightarrow> bool " where - " IsOnes x = ( (x = ((Ones__0 ((make_the_value ((int (size x))) :: ( 'N::len)itself)) :: ( 'N::len)Word.word))))" + " IsOnes x = ( (x = ((Ones__0 ((make_the_value ((int (size x))) )) :: ( 'N::len)Word.word))))" (*val FPMaxNormal : forall 'N . Size 'N => integer -> mword ty1 -> M (mword 'N)*) @@ -4868,35 +4869,35 @@ definition FPMaxNormal :: " int \<Rightarrow>(1)Word.word \<Rightarrow>((regist if (((p00 = (( 16 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 16 :: int)::ii) - (( 5 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 5 bits) = + (let (exp1 :: 5 bits) = ((concat_vec ((Ones__0 ((make_the_value (((( 5 :: int)::ii) - (( 1 :: int)::ii))) :: 4 itself)) :: 4 Word.word)) (vec_of_bits [B0] :: 1 Word.word) :: 5 Word.word)) in (let (frac :: 10 bits) = ((Ones__0 ((make_the_value F :: 10 itself)) :: 10 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 6 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 6 Word.word)) frac :: 16 Word.word)) :: ( 'N::len)Word.word)))))) else if (((p00 = (( 32 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 32 :: int)::ii) - (( 8 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 8 bits) = + (let (exp1 :: 8 bits) = ((concat_vec ((Ones__0 ((make_the_value (((( 8 :: int)::ii) - (( 1 :: int)::ii))) :: 7 itself)) :: 7 Word.word)) (vec_of_bits [B0] :: 1 Word.word) :: 8 Word.word)) in (let (frac :: 23 bits) = ((Ones__0 ((make_the_value F :: 23 itself)) :: 23 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 9 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 9 Word.word)) frac :: 32 Word.word)) :: ( 'N::len)Word.word)))))) else if (((p00 = (( 64 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 64 :: int)::ii) - (( 11 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 11 bits) = + (let (exp1 :: 11 bits) = ((concat_vec ((Ones__0 ((make_the_value (((( 11 :: int)::ii) - (( 1 :: int)::ii))) :: 10 itself)) :: 10 Word.word)) (vec_of_bits [B0] :: 1 Word.word) :: 11 Word.word)) in (let (frac :: 52 bits) = ((Ones__0 ((make_the_value F :: 52 itself)) :: 52 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 12 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 12 Word.word)) frac :: 64 Word.word)) :: ( 'N::len)Word.word)))))) else assert_exp False (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> exit0 () ))" @@ -4909,23 +4910,23 @@ definition FPInfinity :: " int \<Rightarrow>(1)Word.word \<Rightarrow>((registe if (((p00 = (( 16 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 16 :: int)::ii) - (( 5 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 5 bits) = ((Ones__0 ((make_the_value (( 5 :: int)::ii) :: 5 itself)) :: 5 Word.word)) in + (let (exp1 :: 5 bits) = ((Ones__0 ((make_the_value (( 5 :: int)::ii) :: 5 itself)) :: 5 Word.word)) in (let (frac :: 10 bits) = ((Zeros__0 ((make_the_value F :: 10 itself)) :: 10 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 6 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 6 Word.word)) frac :: 16 Word.word)) :: ( 'N::len)Word.word)))))) else if (((p00 = (( 32 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 32 :: int)::ii) - (( 8 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 8 bits) = ((Ones__0 ((make_the_value (( 8 :: int)::ii) :: 8 itself)) :: 8 Word.word)) in + (let (exp1 :: 8 bits) = ((Ones__0 ((make_the_value (( 8 :: int)::ii) :: 8 itself)) :: 8 Word.word)) in (let (frac :: 23 bits) = ((Zeros__0 ((make_the_value F :: 23 itself)) :: 23 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 9 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 9 Word.word)) frac :: 32 Word.word)) :: ( 'N::len)Word.word)))))) else if (((p00 = (( 64 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 64 :: int)::ii) - (( 11 :: int)::ii))) - (( 1 :: int)::ii)) in - (let (exp :: 11 bits) = ((Ones__0 ((make_the_value (( 11 :: int)::ii) :: 11 itself)) :: 11 Word.word)) in + (let (exp1 :: 11 bits) = ((Ones__0 ((make_the_value (( 11 :: int)::ii) :: 11 itself)) :: 11 Word.word)) in (let (frac :: 52 bits) = ((Zeros__0 ((make_the_value F :: 52 itself)) :: 52 Word.word)) in - return ((Word.ucast ((concat_vec ((concat_vec sign exp :: 12 Word.word)) frac :: ( 'N::len)Word.word)) + return ((Word.ucast ((concat_vec ((concat_vec sign exp1 :: 12 Word.word)) frac :: 64 Word.word)) :: ( 'N::len)Word.word)))))) else assert_exp False (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> exit0 () ))" @@ -4939,40 +4940,40 @@ definition FPDefaultNaN :: " int \<Rightarrow> unit \<Rightarrow>((register_val assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 16 :: int)::ii) - (( 5 :: int)::ii))) - (( 1 :: int)::ii)) in (let (sign :: 1 bits) = ((vec_of_bits [B0] :: 1 Word.word)) in - (let (exp :: 5 bits) = ((Ones__0 ((make_the_value (( 5 :: int)::ii) :: 5 itself)) :: 5 Word.word)) in + (let (exp1 :: 5 bits) = ((Ones__0 ((make_the_value (( 5 :: int)::ii) :: 5 itself)) :: 5 Word.word)) in (let (frac :: 10 bits) = ((concat_vec (vec_of_bits [B1] :: 1 Word.word) ((Zeros__0 ((make_the_value ((F - (( 1 :: int)::ii))) :: 9 itself)) :: 9 Word.word)) :: 10 Word.word)) in return ((Word.ucast - ((concat_vec ((concat_vec (vec_of_bits [B0] :: 1 Word.word) exp :: 6 Word.word)) frac - :: ( 'N::len)Word.word)) + ((concat_vec ((concat_vec (vec_of_bits [B0] :: 1 Word.word) exp1 :: 6 Word.word)) frac + :: 16 Word.word)) :: ( 'N::len)Word.word))))))) else if (((p00 = (( 32 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 32 :: int)::ii) - (( 8 :: int)::ii))) - (( 1 :: int)::ii)) in (let (sign :: 1 bits) = ((vec_of_bits [B0] :: 1 Word.word)) in - (let (exp :: 8 bits) = ((Ones__0 ((make_the_value (( 8 :: int)::ii) :: 8 itself)) :: 8 Word.word)) in + (let (exp1 :: 8 bits) = ((Ones__0 ((make_the_value (( 8 :: int)::ii) :: 8 itself)) :: 8 Word.word)) in (let (frac :: 23 bits) = ((concat_vec (vec_of_bits [B1] :: 1 Word.word) ((Zeros__0 ((make_the_value ((F - (( 1 :: int)::ii))) :: 22 itself)) :: 22 Word.word)) :: 23 Word.word)) in return ((Word.ucast - ((concat_vec ((concat_vec (vec_of_bits [B0] :: 1 Word.word) exp :: 9 Word.word)) frac - :: ( 'N::len)Word.word)) + ((concat_vec ((concat_vec (vec_of_bits [B0] :: 1 Word.word) exp1 :: 9 Word.word)) frac + :: 32 Word.word)) :: ( 'N::len)Word.word))))))) else if (((p00 = (( 64 :: int)::ii)))) then assert_exp True (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> ((let (F :: int) = ((((( 64 :: int)::ii) - (( 11 :: int)::ii))) - (( 1 :: int)::ii)) in (let (sign :: 1 bits) = ((vec_of_bits [B0] :: 1 Word.word)) in - (let (exp :: 11 bits) = ((Ones__0 ((make_the_value (( 11 :: int)::ii) :: 11 itself)) :: 11 Word.word)) in + (let (exp1 :: 11 bits) = ((Ones__0 ((make_the_value (( 11 :: int)::ii) :: 11 itself)) :: 11 Word.word)) in (let (frac :: 52 bits) = ((concat_vec (vec_of_bits [B1] :: 1 Word.word) ((Zeros__0 ((make_the_value ((F - (( 1 :: int)::ii))) :: 51 itself)) :: 51 Word.word)) :: 52 Word.word)) in return ((Word.ucast - ((concat_vec ((concat_vec (vec_of_bits [B0] :: 1 Word.word) exp :: 12 Word.word)) frac - :: ( 'N::len)Word.word)) + ((concat_vec ((concat_vec (vec_of_bits [B0] :: 1 Word.word) exp1 :: 12 Word.word)) frac + :: 64 Word.word)) :: ( 'N::len)Word.word))))))) else assert_exp False (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> exit0 () ))" @@ -4983,7 +4984,7 @@ definition FPConvertNaN :: " int \<Rightarrow>('N::len)Word.word \<Rightarrow>( " FPConvertNaN (M__tv :: int) op1 = ( ((assert_exp ((((((((int (size op1))) = (( 16 :: int)::ii)))) \<or> ((((((((int (size op1))) = (( 32 :: int)::ii)))) \<or> (((((int (size op1))) = (( 64 :: int)::ii)))))))))) (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> assert_exp ((((((M__tv = (( 16 :: int)::ii)))) \<or> ((((((M__tv = (( 32 :: int)::ii)))) \<or> (((M__tv = (( 64 :: int)::ii)))))))))) (''((M == 16) || ((M == 32) || (M == 64)))'')) \<then> - (undefined_bitvector M__tv :: (( 'M::len)Word.word) M)) \<bind> (\<lambda> (result :: 'M bits) . + (undefined_bitvector M__tv :: (( 'M::len)Word.word) M)) \<bind> (\<lambda> result . (undefined_bitvector (( 51 :: int)::ii) :: ( 51 Word.word) M) \<bind> (\<lambda> (frac :: 51 bits) . (let (sign :: 1 bits) = ((vec_of_bits [access_vec_dec op1 ((((int (size op1))) - (( 1 :: int)::ii)))] :: 1 Word.word)) in @@ -5003,24 +5004,30 @@ definition FPConvertNaN :: " int \<Rightarrow>('N::len)Word.word \<Rightarrow>( ((Zeros__0 ((make_the_value (( 42 :: int)::ii) :: 42 itself)) :: 42 Word.word)) :: 51 Word.word))) in (let p00 = (int (size result)) in - (let (result :: 'M bits) = + (let result = (if (((p00 = (( 64 :: int)::ii)))) then - (concat_vec - ((concat_vec sign - ((Ones__0 ((make_the_value (((( 64 :: int)::ii) - (( 52 :: int)::ii))) )) :: 12 Word.word)) - :: 13 Word.word)) frac + (Word.ucast + ((concat_vec + ((concat_vec sign + ((Ones__0 ((make_the_value (((( 64 :: int)::ii) - (( 52 :: int)::ii))) )) :: 12 Word.word)) + :: 13 Word.word)) frac + :: 64 Word.word)) :: ( 'M::len)Word.word) else if (((p00 = (( 32 :: int)::ii)))) then - (concat_vec - ((concat_vec sign - ((Ones__0 ((make_the_value (((( 32 :: int)::ii) - (( 23 :: int)::ii))) )) :: 9 Word.word)) - :: 10 Word.word)) ((slice frac (( 29 :: int)::ii) (( 22 :: int)::ii) :: 22 Word.word)) + (Word.ucast + ((concat_vec + ((concat_vec sign + ((Ones__0 ((make_the_value (((( 32 :: int)::ii) - (( 23 :: int)::ii))) )) :: 9 Word.word)) + :: 10 Word.word)) ((slice frac (( 29 :: int)::ii) (( 22 :: int)::ii) :: 22 Word.word)) + :: 32 Word.word)) :: ( 'M::len)Word.word) else - (concat_vec - ((concat_vec sign - ((Ones__0 ((make_the_value ((p00 - (( 10 :: int)::ii))) )) :: 6 Word.word)) - :: 7 Word.word)) ((slice frac (( 42 :: int)::ii) (( 9 :: int)::ii) :: 9 Word.word)) + (Word.ucast + ((concat_vec + ((concat_vec sign + ((Ones__0 ((make_the_value ((p00 - (( 10 :: int)::ii))) )) :: 6 Word.word)) + :: 7 Word.word)) ((slice frac (( 42 :: int)::ii) (( 9 :: int)::ii) :: 9 Word.word)) + :: 16 Word.word)) :: ( 'M::len)Word.word)) in return result))))))))" @@ -5287,19 +5294,19 @@ definition ComputePAC :: "(64)Word.word \<Rightarrow>(64)Word.word \<Rightarrow (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (roundkey :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (modk0 :: 64 bits) . (hex_slice (''0xC0AC29B7C97C50DD'') (( 64 :: int)::ii) (( 0 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (Alpha :: 64 bits) . - read_reg RC_ref \<bind> (\<lambda> (w__0 :: ( 64 Word.word) list) . + read_reg RC_ref \<bind> (\<lambda> (w__0 :: ( 64 bits) list) . (hex_slice (''0x0'') (( 64 :: int)::ii) (( 0 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (write_reg RC_ref ((update_list_dec w__0 (( 0 :: int)::ii) w__1 :: ( 64 Word.word) list)) \<then> - read_reg RC_ref) \<bind> (\<lambda> (w__2 :: ( 64 Word.word) list) . + read_reg RC_ref) \<bind> (\<lambda> (w__2 :: ( 64 bits) list) . (hex_slice (''0x13198A2E03707344'') (( 64 :: int)::ii) (( 0 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . (write_reg RC_ref ((update_list_dec w__2 (( 1 :: int)::ii) w__3 :: ( 64 Word.word) list)) \<then> - read_reg RC_ref) \<bind> (\<lambda> (w__4 :: ( 64 Word.word) list) . + read_reg RC_ref) \<bind> (\<lambda> (w__4 :: ( 64 bits) list) . (hex_slice (''0xA493822299F31D0'') (( 64 :: int)::ii) (( 0 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__5 :: 64 Word.word) . (write_reg RC_ref ((update_list_dec w__4 (( 2 :: int)::ii) w__5 :: ( 64 Word.word) list)) \<then> - read_reg RC_ref) \<bind> (\<lambda> (w__6 :: ( 64 Word.word) list) . + read_reg RC_ref) \<bind> (\<lambda> (w__6 :: ( 64 bits) list) . (hex_slice (''0x82EFA98EC4E6C89'') (( 64 :: int)::ii) (( 0 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 Word.word) . (write_reg RC_ref ((update_list_dec w__6 (( 3 :: int)::ii) w__7 :: ( 64 Word.word) list)) \<then> - read_reg RC_ref) \<bind> (\<lambda> (w__8 :: ( 64 Word.word) list) . + read_reg RC_ref) \<bind> (\<lambda> (w__8 :: ( 64 bits) list) . (hex_slice (''0x452821E638D01377'') (( 64 :: int)::ii) (( 0 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__9 :: 64 Word.word) . write_reg RC_ref ((update_list_dec w__8 (( 4 :: int)::ii) w__9 :: ( 64 Word.word) list)) \<then> ((let modk0 = @@ -5320,51 +5327,51 @@ definition ComputePAC :: "(64)Word.word \<Rightarrow>(64)Word.word \<Rightarrow read_reg RC_ref \<bind> (\<lambda> (w__10 :: ( 64 bits) list) . (let workingval = ((xor_vec workingval ((access_list_dec w__10 i :: 64 Word.word)) :: 64 Word.word)) in (if ((i > (( 0 :: int)::ii))) then - (PACCellShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__11 :: 64 bits) . + (PACCellShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__11 :: 64 Word.word) . (let workingval = w__11 in (PACMult workingval :: ( 64 Word.word) M))) else return workingval) \<bind> (\<lambda> (workingval :: 64 bits) . - (PACSub workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__13 :: 64 bits) . + (PACSub workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__13 :: 64 Word.word) . (let workingval = w__13 in - (TweakShuffle ((slice runningmod (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: 64 - bits) . + (TweakShuffle ((slice runningmod (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: + 64 Word.word) . (let (runningmod :: 64 bits) = w__14 in return (roundkey, runningmod, workingval))))))))))))) \<bind> (\<lambda> varstup . (let ((roundkey :: 64 bits), (runningmod :: 64 bits), (workingval :: 64 bits)) = varstup in (let roundkey = ((xor_vec modk0 runningmod :: 64 Word.word)) in (let workingval = ((xor_vec workingval roundkey :: 64 Word.word)) in - (PACCellShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 bits) . + (PACCellShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 Word.word) . (let workingval = w__15 in - (PACMult workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__16 :: 64 bits) . + (PACMult workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__16 :: 64 Word.word) . (let workingval = w__16 in - (PACSub workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: 64 bits) . + (PACSub workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: 64 Word.word) . (let workingval = w__17 in - (PACCellShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__18 :: 64 bits) . + (PACCellShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__18 :: 64 Word.word) . (let workingval = w__18 in - (PACMult workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__19 :: 64 bits) . + (PACMult workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__19 :: 64 Word.word) . (let workingval = w__19 in (let workingval = ((xor_vec key1 workingval :: 64 Word.word)) in - (PACCellInvShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__20 :: 64 bits) . + (PACCellInvShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__20 :: 64 Word.word) . (let workingval = w__20 in - (PACInvSub workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__21 :: 64 bits) . + (PACInvSub workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__21 :: 64 Word.word) . (let workingval = w__21 in - (PACMult workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__22 :: 64 bits) . + (PACMult workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__22 :: 64 Word.word) . (let workingval = w__22 in - (PACCellInvShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__23 :: 64 bits) . + (PACCellInvShuffle workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__23 :: 64 Word.word) . (let workingval = w__23 in (let workingval = ((xor_vec workingval key0 :: 64 Word.word)) in (let workingval = ((xor_vec workingval runningmod :: 64 Word.word)) in (foreachM (index_list (( 0 :: int)::ii) (( 4 :: int)::ii) (( 1 :: int)::ii)) (roundkey, runningmod, workingval) (\<lambda> i varstup . (let (roundkey, runningmod, workingval) = varstup in - (PACInvSub workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__24 :: 64 bits) . + (PACInvSub workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__24 :: 64 Word.word) . (let workingval = w__24 in (if ((i < (( 4 :: int)::ii))) then - (PACMult workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__25 :: 64 bits) . + (PACMult workingval :: ( 64 Word.word) M) \<bind> (\<lambda> (w__25 :: 64 Word.word) . (let workingval = w__25 in (PACCellInvShuffle workingval :: ( 64 Word.word) M))) else return workingval) \<bind> (\<lambda> (workingval :: 64 bits) . - (TweakInvShuffle ((slice runningmod (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__27 :: 64 - bits) . + (TweakInvShuffle ((slice runningmod (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__27 :: + 64 Word.word) . (let runningmod = w__27 in (let roundkey = ((xor_vec key1 runningmod :: 64 Word.word)) in read_reg RC_ref \<bind> (\<lambda> (w__28 :: ( 64 bits) list) . @@ -5389,8 +5396,7 @@ definition Align__0 :: " int \<Rightarrow> int \<Rightarrow> int " where definition Align__1 :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>('N::len)Word.word " where " Align__1 x y = ( - (GetSlice_int ((make_the_value ((int (size x))) :: ( 'N::len)itself)) ((Align__0 ((Word.uint x)) y)) (( 0 :: int)::ii) - :: ( 'N::len)Word.word))" + (GetSlice_int ((make_the_value ((int (size x))) )) ((Align__0 ((Word.uint x)) y)) (( 0 :: int)::ii) :: ( 'N::len)Word.word))" (*val aset__Mem : forall 'p8_times_size_ . Size 'p8_times_size_ => AddressDescriptor -> integer -> AccessDescriptor -> mword 'p8_times_size_ -> M unit*) @@ -5403,11 +5409,11 @@ definition aset__Mem :: " AddressDescriptor \<Rightarrow> int \<Rightarrow> Acc (hex_slice (''0x13000000'') (( 52 :: int)::ii) (( 0 :: int)::ii) :: ( 52 Word.word) M)) \<bind> (\<lambda> (w__0 :: 52 Word.word) . if (((address = w__0))) then if (((((Word.uint value_name)) = (( 4 :: int)::ii)))) then - (let (_ :: unit) = (prerr ([(CHR ''P''), (CHR ''r''), (CHR ''o''), (CHR ''g''), (CHR ''r''), (CHR ''a''), (CHR ''m''), (CHR '' ''), (CHR ''e''), (CHR ''x''), (CHR ''i''), (CHR ''t''), (CHR ''e''), (CHR ''d''), (CHR '' ''), (CHR ''b''), (CHR ''y''), (CHR '' ''), (CHR ''w''), (CHR ''r''), (CHR ''i''), (CHR ''t''), (CHR ''i''), (CHR ''n''), (CHR ''g''), (CHR '' ''), (CHR ''^''), (CHR ''D''), (CHR '' ''), (CHR ''t''), (CHR ''o''), (CHR '' ''), (CHR ''T''), (CHR ''U''), (CHR ''B''), (CHR ''E''), (char_of_nat 10)])) in + (let (_ :: unit) = (prerr ([(CHR ''P''), (CHR ''r''), (CHR ''o''), (CHR ''g''), (CHR ''r''), (CHR ''a''), (CHR ''m''), (CHR '' ''), (CHR ''e''), (CHR ''x''), (CHR ''i''), (CHR ''t''), (CHR ''e''), (CHR ''d''), (CHR '' ''), (CHR ''b''), (CHR ''y''), (CHR '' ''), (CHR ''w''), (CHR ''r''), (CHR ''i''), (CHR ''t''), (CHR ''i''), (CHR ''n''), (CHR ''g''), (CHR '' ''), (CHR ''^''), (CHR ''D''), (CHR '' ''), (CHR ''t''), (CHR ''o''), (CHR '' ''), (CHR ''T''), (CHR ''U''), (CHR ''B''), (CHR ''E''), (CHR 0x27)])) in exit0 () ) else return ((putchar ((Word.uint ((slice value_name (( 0 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)))))) else - (read_reg Memory_ref :: ( 52 Word.word) M) \<bind> (\<lambda> (w__1 :: 52 Word.word) . + (read_reg Memory_ref :: ( 52 Word.word) M) \<bind> (\<lambda> (w__1 :: 52 bits) . WriteRAM ((make_the_value (( 52 :: int)::ii) :: 52 itself)) size1 w__1 address value_name)))))" @@ -5418,7 +5424,7 @@ definition aget__Mem :: " AddressDescriptor \<Rightarrow> int \<Rightarrow> Acc assert_exp ((((((size1 = (( 1 :: int)::ii)))) \<or> ((((((size1 = (( 2 :: int)::ii)))) \<or> ((((((size1 = (( 4 :: int)::ii)))) \<or> ((((((size1 = (( 8 :: int)::ii)))) \<or> (((size1 = (( 16 :: int)::ii)))))))))))))))) (''((size == 1) || ((size == 2) || ((size == 4) || ((size == 8) || (size == 16)))))'') \<then> ((let (address :: 52 bits) = ((FullAddress_physicaladdress (AddressDescriptor_paddress desc))) in (assert_exp (((address = ((Align__1 address size1 :: 52 Word.word))))) (''(address == Align(address, size))'') \<then> - (read_reg Memory_ref :: ( 52 Word.word) M)) \<bind> (\<lambda> (w__0 :: 52 Word.word) . + (read_reg Memory_ref :: ( 52 Word.word) M)) \<bind> (\<lambda> (w__0 :: 52 bits) . (ReadRAM ((make_the_value (( 52 :: int)::ii) :: 52 itself)) size1 w__0 address :: (( 'p8_times_size_::len)Word.word) M)))))" @@ -5430,7 +5436,7 @@ definition aset_X :: " int \<Rightarrow>('width::len)Word.word \<Rightarrow>((r (assert_exp (((((n \<ge> (( 0 :: int)::ii))) \<and> ((n \<le> (( 31 :: int)::ii)))))) (''((n >= 0) && (n <= 31))'') \<then> assert_exp ((((((((int (size value_name))) = (( 32 :: int)::ii)))) \<or> (((((int (size value_name))) = (( 64 :: int)::ii))))))) (''((width == 32) || (width == 64))'')) \<then> (if (((n \<noteq> (( 31 :: int)::ii)))) then - read_reg R_ref \<bind> (\<lambda> (w__0 :: ( 64 Word.word) list) . + read_reg R_ref \<bind> (\<lambda> (w__0 :: ( 64 bits) list) . (ZeroExtend__0 value_name ((make_the_value (( 64 :: int)::ii) :: 64 itself)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . write_reg R_ref ((update_list_dec w__0 n w__1 :: ( 64 Word.word) list)))) @@ -5478,7 +5484,7 @@ definition AArch64_ResetGeneralRegisters :: " unit \<Rightarrow>((register_valu " AArch64_ResetGeneralRegisters _ = ( (foreachM (index_list (( 0 :: int)::ii) (( 30 :: int)::ii) (( 1 :: int)::ii)) () (\<lambda> i unit_var . - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . aset_X i w__0))))" + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . aset_X i w__0))))" (*val aset_ELR__0 : mword ty2 -> mword ty64 -> M unit*) @@ -5511,7 +5517,7 @@ definition aget_X :: " int \<Rightarrow> int \<Rightarrow>((register_value),((' (if (((n \<noteq> (( 31 :: int)::ii)))) then read_reg R_ref \<bind> (\<lambda> (w__0 :: ( 64 bits) list) . return ((slice ((access_list_dec w__0 n :: 64 Word.word)) (( 0 :: int)::ii) width__tv :: ( 'width::len)Word.word))) - else return ((Zeros__0 ((make_the_value width__tv :: ( 'width::len)itself)) :: ( 'width::len)Word.word))))" + else return ((Zeros__0 ((make_the_value width__tv )) :: ( 'width::len)Word.word))))" (*val aarch64_system_sysops : bool -> ii -> ii -> ii -> ii -> ii -> ii -> M unit*) @@ -6783,7 +6789,7 @@ definition integer_arithmetic_addsub_carry_decode :: "(1)Word.word \<Rightarrow definition ExtendReg :: " int \<Rightarrow> int \<Rightarrow> ExtendType \<Rightarrow> int \<Rightarrow>((register_value),(('N::len)Word.word),(exception))monad " where " ExtendReg (N__tv :: int) reg typ1 shift = ( (assert_exp (((((shift \<ge> (( 0 :: int)::ii))) \<and> ((shift \<le> (( 4 :: int)::ii)))))) (''((shift >= 0) && (shift <= 4))'') \<then> - (aget_X N__tv reg :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> (val_name :: 'N bits) . + (aget_X N__tv reg :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> val_name . undefined_bool () \<bind> (\<lambda> (unsigned :: bool) . undefined_int () \<bind> (\<lambda> (len :: ii) . (let (len :: ii) = @@ -6852,7 +6858,7 @@ definition ROR_C :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>((regist ((let (m :: ii) = (shift mod ((int (size x)))) in (LSR x m :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'N::len)Word.word) . (LSL x ((((int (size x))) - ((ex_int m)))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: ( 'N::len)Word.word) . - (let (result :: 'N bits) = ((or_vec w__0 w__1 :: ( 'N::len)Word.word)) in + (let result = ((or_vec w__0 w__1 :: ( 'N::len)Word.word)) in (let (carry_out :: 1 bits) = ((vec_of_bits [access_vec_dec result ((((int (size result))) - (( 1 :: int)::ii)))] :: 1 Word.word)) in return (result, carry_out))))))))" @@ -6864,11 +6870,11 @@ definition ROR :: "('N::len)Word.word \<Rightarrow> int \<Rightarrow>((register " ROR x shift = ( (assert_exp ((shift \<ge> (( 0 :: int)::ii))) (''(shift >= 0)'') \<then> (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (anon10 :: 1 bits) . - (undefined_bitvector ((int (size x))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size x))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . if (((shift = (( 0 :: int)::ii)))) then return x else - (ROR_C x shift :: ((( 'N::len)Word.word * 1 Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in - (let (result :: 'N bits) = tup__0 in + (ROR_C x shift ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (let result = tup__0 in (let (anon10 :: 1 bits) = tup__1 in return result)))))))" @@ -6882,17 +6888,17 @@ definition aarch64_integer_bitfield :: " int \<Rightarrow> int \<Rightarrow> in ((assert_exp True (''datasize constraint'') \<then> assert_exp True (''dbytes constraint'')) \<then> (if inzero then return ((Zeros__1 datasize () :: ( 'datasize::len)Word.word)) - else (aget_X datasize d :: (( 'datasize::len)Word.word) M))) \<bind> (\<lambda> (dst :: 'datasize bits) . - (aget_X datasize n :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (src :: 'datasize bits) . + else (aget_X datasize d :: (( 'datasize::len)Word.word) M))) \<bind> (\<lambda> dst . + (aget_X datasize n :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> src . (ROR src R1 :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__1 :: ( 'datasize::len)Word.word) . - (let (bot :: 'datasize bits) = + (let bot = ((or_vec ((and_vec dst ((not_vec wmask :: ( 'datasize::len)Word.word)) :: ( 'datasize::len)Word.word)) ((and_vec w__1 wmask :: ( 'datasize::len)Word.word)) :: ( 'datasize::len)Word.word)) in (if extend1 then (Replicate ((int (size bot))) (vec_of_bits [access_vec_dec src S] :: 1 Word.word) :: (( 'datasize::len)Word.word) M) - else return dst) \<bind> (\<lambda> (top1 :: 'datasize bits) . + else return dst) \<bind> (\<lambda> top1 . aset_X d ((or_vec ((and_vec top1 ((not_vec tmask :: ( 'datasize::len)Word.word)) :: ( 'datasize::len)Word.word)) ((and_vec bot tmask :: ( 'datasize::len)Word.word)) @@ -6903,7 +6909,7 @@ definition aarch64_integer_bitfield :: " int \<Rightarrow> int \<Rightarrow> in definition ShiftReg :: " int \<Rightarrow> int \<Rightarrow> ShiftType \<Rightarrow> int \<Rightarrow>((register_value),(('N::len)Word.word),(exception))monad " where " ShiftReg (N__tv :: int) reg typ1 amount = ( - (aget_X N__tv reg :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (aget_X N__tv reg :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . (case typ1 of ShiftType_LSL => (LSL result amount :: (( 'N::len)Word.word) M) | ShiftType_LSR => (LSR result amount :: (( 'N::len)Word.word) M) @@ -6922,8 +6928,8 @@ definition aarch64_integer_shift_variable :: " int \<Rightarrow> int \<Rightarr assert_exp True (''dbytes constraint'')) \<then> (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M)) \<bind> (\<lambda> (result :: 8 bits) . (aget_X (( 8 :: int)::ii) m :: ( 8 Word.word) M) \<bind> (\<lambda> (operand2 :: 8 bits) . - (ShiftReg (( 8 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 8 :: int)::ii))) :: ( 8 Word.word) M) \<bind> (\<lambda> (w__0 :: 8 - bits) . + (ShiftReg (( 8 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 8 :: int)::ii))) :: ( 8 Word.word) M) \<bind> (\<lambda> (w__0 :: + 8 Word.word) . (let result = w__0 in aset_X d result))))) else if (((l__173 = (( 16 :: int)::ii)))) then @@ -6932,8 +6938,8 @@ definition aarch64_integer_shift_variable :: " int \<Rightarrow> int \<Rightarr assert_exp True (''dbytes constraint'')) \<then> (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M)) \<bind> (\<lambda> (result :: 16 bits) . (aget_X (( 16 :: int)::ii) m :: ( 16 Word.word) M) \<bind> (\<lambda> (operand2 :: 16 bits) . - (ShiftReg (( 16 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 16 :: int)::ii))) :: ( 16 Word.word) M) \<bind> (\<lambda> (w__1 :: 16 - bits) . + (ShiftReg (( 16 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 16 :: int)::ii))) :: ( 16 Word.word) M) \<bind> (\<lambda> (w__1 :: + 16 Word.word) . (let result = w__1 in aset_X d result))))) else if (((l__173 = (( 32 :: int)::ii)))) then @@ -6942,8 +6948,8 @@ definition aarch64_integer_shift_variable :: " int \<Rightarrow> int \<Rightarr assert_exp True (''dbytes constraint'')) \<then> (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (result :: 32 bits) . (aget_X (( 32 :: int)::ii) m :: ( 32 Word.word) M) \<bind> (\<lambda> (operand2 :: 32 bits) . - (ShiftReg (( 32 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 32 :: int)::ii))) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 - bits) . + (ShiftReg (( 32 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 32 :: int)::ii))) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: + 32 Word.word) . (let result = w__2 in aset_X d result))))) else if (((l__173 = (( 64 :: int)::ii)))) then @@ -6952,8 +6958,8 @@ definition aarch64_integer_shift_variable :: " int \<Rightarrow> int \<Rightarr assert_exp True (''dbytes constraint'')) \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (result :: 64 bits) . (aget_X (( 64 :: int)::ii) m :: ( 64 Word.word) M) \<bind> (\<lambda> (operand2 :: 64 bits) . - (ShiftReg (( 64 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 64 :: int)::ii))) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 - bits) . + (ShiftReg (( 64 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 64 :: int)::ii))) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: + 64 Word.word) . (let result = w__3 in aset_X d result))))) else if (((l__173 = (( 128 :: int)::ii)))) then @@ -6962,8 +6968,8 @@ definition aarch64_integer_shift_variable :: " int \<Rightarrow> int \<Rightarr assert_exp True (''dbytes constraint'')) \<then> (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M)) \<bind> (\<lambda> (result :: 128 bits) . (aget_X (( 128 :: int)::ii) m :: ( 128 Word.word) M) \<bind> (\<lambda> (operand2 :: 128 bits) . - (ShiftReg (( 128 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 128 :: int)::ii))) :: ( 128 Word.word) M) \<bind> (\<lambda> (w__4 :: 128 - bits) . + (ShiftReg (( 128 :: int)::ii) n shift_type ((((Word.uint operand2)) mod (( 128 :: int)::ii))) :: ( 128 Word.word) M) \<bind> (\<lambda> (w__4 :: + 128 Word.word) . (let result = w__4 in aset_X d result))))) else @@ -7529,10 +7535,10 @@ definition aarch64_integer_logical_immediate :: " int \<Rightarrow>('datasize:: (let dbytes = (ex_int ((datasize div (( 8 :: int)::ii)))) in ((assert_exp True (''datasize constraint'') \<then> assert_exp True (''dbytes constraint'')) \<then> - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> (result :: 'datasize bits) . - (aget_X datasize n :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (operand1 :: 'datasize bits) . - (let (operand2 :: 'datasize bits) = imm in - (let (result :: 'datasize bits) = + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> result . + (aget_X datasize n :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> operand1 . + (let operand2 = imm in + (let result = ((case op1 of LogicalOp_AND => (and_vec operand1 operand2 :: ( 'datasize::len)Word.word) | LogicalOp_ORR => (or_vec operand1 operand2 :: ( 'datasize::len)Word.word) @@ -7572,21 +7578,21 @@ definition aarch64_integer_arithmetic_addsub_immediate :: " int \<Rightarrow>(' (let dbytes = (ex_int ((datasize div (( 8 :: int)::ii)))) in ((assert_exp True (''datasize constraint'') \<then> assert_exp True (''dbytes constraint'')) \<then> - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> (result :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> result . (if (((n = (( 31 :: int)::ii)))) then (aget_SP datasize () :: (( 'datasize::len)Word.word) M) - else (aget_X datasize n :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> (operand1 :: 'datasize bits) . - (let (operand2 :: 'datasize bits) = imm in + else (aget_X datasize n :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> operand1 . + (let operand2 = imm in (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (nzcv :: 4 bits) . (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (carry_in :: 1 bits) . - (let ((carry_in :: 1 bits), (operand2 :: 'datasize bits)) = + (let ((carry_in :: 1 bits), operand2) = (if sub_op then - (let (operand2 :: 'datasize bits) = ((not_vec operand2 :: ( 'datasize::len)Word.word)) in + (let operand2 = ((not_vec operand2 :: ( 'datasize::len)Word.word)) in (let (carry_in :: 1 bits) = ((vec_of_bits [B1] :: 1 Word.word)) in (carry_in, operand2))) else (let (carry_in :: 1 bits) = ((vec_of_bits [B0] :: 1 Word.word)) in (carry_in, operand2))) in - (let (tup__0, tup__1) = ((AddWithCarry operand1 operand2 carry_in :: (( 'datasize::len)Word.word * 4 Word.word))) in + (let (tup__0, tup__1) = (AddWithCarry operand1 operand2 carry_in ) in (let result = tup__0 in (let nzcv = tup__1 in (if setflags then @@ -8044,7 +8050,7 @@ definition Auth :: "(64)Word.word \<Rightarrow>(64)Word.word \<Rightarrow>(128) :: 64 Word.word)) in (ComputePAC original_ptr modifier ((subrange_vec_dec K (( 127 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) ((subrange_vec_dec K (( 63 :: int)::ii) (( 0 :: int)::ii) :: 64 Word.word)) - :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . + :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (let (PAC :: 64 bits) = w__1 in (let (result :: 64 bits) = (if tbi then @@ -8213,12 +8219,12 @@ definition FPProcessException :: " FPExc \<Rightarrow>(32)Word.word \<Rightarro else UsingAArch32 () \<bind> (\<lambda> (w__0 :: bool) . if w__0 then - (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . + (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . write_reg FPSCR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__1 cumul (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word))) else - (read_reg FPSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg FPSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . write_reg FPSR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__2 cumul (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word))))))))" @@ -8270,16 +8276,16 @@ definition FPRoundBase :: " int \<Rightarrow> real \<Rightarrow>(32)Word.word \ if (((((((((((((vec_of_bits [access_vec_dec fpcr (( 24 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) \<and> False))) \<or> (((((((vec_of_bits [access_vec_dec fpcr (( 19 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) \<and> True)))))) \<and> ((((ex_int exponent)) < ((ex_int minimum_exp))))))) then UsingAArch32 () \<bind> (\<lambda> (w__0 :: bool) . ((if w__0 then - (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . + (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . write_reg FPSCR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__1 (( 3 :: int)::ii) (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word))) else - (read_reg FPSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg FPSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . write_reg FPSR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__2 (( 3 :: int)::ii) (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word)))) \<then> - (FPZero (( 16 :: int)::ii) sign :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> (w__3 :: ( 'N::len)Word.word) . + (FPZero (( 16 :: int)::ii) sign :: ( 16 Word.word) M)) \<bind> (\<lambda> (w__3 :: 16 Word.word) . return ((Word.ucast w__3 :: ( 'N::len)Word.word)))) else (let (biased_exp :: ii) = @@ -8433,16 +8439,16 @@ definition FPRoundBase :: " int \<Rightarrow> real \<Rightarrow>(32)Word.word \ if (((((((((((((vec_of_bits [access_vec_dec fpcr (( 24 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) \<and> True))) \<or> (((((((vec_of_bits [access_vec_dec fpcr (( 19 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) \<and> False)))))) \<and> ((((ex_int exponent)) < ((ex_int minimum_exp))))))) then UsingAArch32 () \<bind> (\<lambda> (w__8 :: bool) . ((if w__8 then - (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 Word.word) . + (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 bits) . write_reg FPSCR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__9 (( 3 :: int)::ii) (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word))) else - (read_reg FPSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . + (read_reg FPSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . write_reg FPSR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__10 (( 3 :: int)::ii) (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word)))) \<then> - (FPZero (( 32 :: int)::ii) sign :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> (w__11 :: ( 'N::len)Word.word) . + (FPZero (( 32 :: int)::ii) sign :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__11 :: 32 Word.word) . return ((Word.ucast w__11 :: ( 'N::len)Word.word)))) else (let (biased_exp :: ii) = @@ -8573,16 +8579,16 @@ definition FPRoundBase :: " int \<Rightarrow> real \<Rightarrow>(32)Word.word \ if (((((((((((((vec_of_bits [access_vec_dec fpcr (( 24 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) \<and> True))) \<or> (((((((vec_of_bits [access_vec_dec fpcr (( 19 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) \<and> False)))))) \<and> ((((ex_int exponent)) < ((ex_int minimum_exp))))))) then UsingAArch32 () \<bind> (\<lambda> (w__16 :: bool) . ((if w__16 then - (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__17 :: 32 Word.word) . + (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__17 :: 32 bits) . write_reg FPSCR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__17 (( 3 :: int)::ii) (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word))) else - (read_reg FPSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 Word.word) . + (read_reg FPSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 bits) . write_reg FPSR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__18 (( 3 :: int)::ii) (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word)))) \<then> - (FPZero (( 64 :: int)::ii) sign :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> (w__19 :: ( 'N::len)Word.word) . + (FPZero (( 64 :: int)::ii) sign :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__19 :: 64 Word.word) . return ((Word.ucast w__19 :: ( 'N::len)Word.word)))) else (let (biased_exp :: ii) = @@ -8706,7 +8712,7 @@ definition FixedToFP :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow> int " FixedToFP (N__tv :: int) op1 fbits unsigned fpcr rounding = ( ((assert_exp ((((((N__tv = (( 16 :: int)::ii)))) \<or> ((((((N__tv = (( 32 :: int)::ii)))) \<or> (((N__tv = (( 64 :: int)::ii)))))))))) ('''') \<then> assert_exp ((((((((int (size op1))) = (( 16 :: int)::ii)))) \<or> ((((((((int (size op1))) = (( 32 :: int)::ii)))) \<or> (((((int (size op1))) = (( 64 :: int)::ii)))))))))) ('''')) \<then> - (undefined_bitvector N__tv :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector N__tv :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> result . (assert_exp ((fbits \<ge> (( 0 :: int)::ii))) ('''') \<then> assert_exp (((rounding \<noteq> FPRounding_ODD))) ('''')) \<then> ((let (int_operand :: ii) = (asl_Int op1 unsigned) in @@ -8735,12 +8741,12 @@ definition FPProcessNaN :: " FPType \<Rightarrow>('N::len)Word.word \<Rightarro else (let (op1 :: 64 Word.word) = ((Word.ucast op1 :: 64 Word.word)) in (( 51 :: int)::ii))) in - (let (result :: 'N bits) = op1 in + (let result = op1 in (if (((typ1 = FPType_SNaN))) then (let result = ((set_slice ((int (size op1))) (( 1 :: int)::ii) result topfrac (vec_of_bits [B1] :: 1 Word.word) :: ( 'N::len)Word.word)) in FPProcessException FPExc_InvalidOp fpcr \<then> return result) - else return result) \<bind> (\<lambda> (result :: 'N bits) . + else return result) \<bind> (\<lambda> result . if ((((vec_of_bits [access_vec_dec fpcr (( 25 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) then (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) else return result))))))" @@ -8751,41 +8757,41 @@ definition FPProcessNaN :: " FPType \<Rightarrow>('N::len)Word.word \<Rightarro definition FPProcessNaNs3 :: " FPType \<Rightarrow> FPType \<Rightarrow> FPType \<Rightarrow>('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Rightarrow>(32)Word.word \<Rightarrow>((register_value),(bool*('N::len)Word.word),(exception))monad " where " FPProcessNaNs3 type1 type2 type3 op1 op2 op3 fpcr = ( (assert_exp ((((((((int (size op1))) = (( 16 :: int)::ii)))) \<or> ((((((((int (size op1))) = (( 32 :: int)::ii)))) \<or> (((((int (size op1))) = (( 64 :: int)::ii)))))))))) (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . if (((type1 = FPType_SNaN))) then (let done1 = True in - (FPProcessNaN type1 op1 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'N bits) . - (let (result :: 'N bits) = w__0 in + (FPProcessNaN type1 op1 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'N::len)Word.word) . + (let result = w__0 in return (done1, result)))) else if (((type2 = FPType_SNaN))) then (let done1 = True in - (FPProcessNaN type2 op2 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: 'N bits) . - (let (result :: 'N bits) = w__1 in + (FPProcessNaN type2 op2 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: ( 'N::len)Word.word) . + (let result = w__1 in return (done1, result)))) else if (((type3 = FPType_SNaN))) then (let done1 = True in - (FPProcessNaN type3 op3 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'N bits) . - (let (result :: 'N bits) = w__2 in + (FPProcessNaN type3 op3 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__2 :: ( 'N::len)Word.word) . + (let result = w__2 in return (done1, result)))) else if (((type1 = FPType_QNaN))) then (let done1 = True in - (FPProcessNaN type1 op1 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__3 :: 'N bits) . - (let (result :: 'N bits) = w__3 in + (FPProcessNaN type1 op1 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'N::len)Word.word) . + (let result = w__3 in return (done1, result)))) else if (((type2 = FPType_QNaN))) then (let done1 = True in - (FPProcessNaN type2 op2 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__4 :: 'N bits) . - (let (result :: 'N bits) = w__4 in + (FPProcessNaN type2 op2 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__4 :: ( 'N::len)Word.word) . + (let result = w__4 in return (done1, result)))) else if (((type3 = FPType_QNaN))) then (let done1 = True in - (FPProcessNaN type3 op3 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__5 :: 'N bits) . - (let (result :: 'N bits) = w__5 in + (FPProcessNaN type3 op3 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__5 :: ( 'N::len)Word.word) . + (let result = w__5 in return (done1, result)))) else (let (done1 :: bool) = False in - (let (result :: 'N bits) = ((Zeros__1 ((int (size op1))) () :: ( 'N::len)Word.word)) in + (let result = ((Zeros__1 ((int (size op1))) () :: ( 'N::len)Word.word)) in return (done1, result))))))" @@ -8794,31 +8800,31 @@ definition FPProcessNaNs3 :: " FPType \<Rightarrow> FPType \<Rightarrow> FPType definition FPProcessNaNs :: " FPType \<Rightarrow> FPType \<Rightarrow>('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Rightarrow>(32)Word.word \<Rightarrow>((register_value),(bool*('N::len)Word.word),(exception))monad " where " FPProcessNaNs type1 type2 op1 op2 fpcr = ( (assert_exp ((((((((int (size op1))) = (( 16 :: int)::ii)))) \<or> ((((((((int (size op1))) = (( 32 :: int)::ii)))) \<or> (((((int (size op1))) = (( 64 :: int)::ii)))))))))) (''((N == 16) || ((N == 32) || (N == 64)))'') \<then> - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . if (((type1 = FPType_SNaN))) then (let done1 = True in - (FPProcessNaN type1 op1 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'N bits) . - (let (result :: 'N bits) = w__0 in + (FPProcessNaN type1 op1 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'N::len)Word.word) . + (let result = w__0 in return (done1, result)))) else if (((type2 = FPType_SNaN))) then (let done1 = True in - (FPProcessNaN type2 op2 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: 'N bits) . - (let (result :: 'N bits) = w__1 in + (FPProcessNaN type2 op2 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: ( 'N::len)Word.word) . + (let result = w__1 in return (done1, result)))) else if (((type1 = FPType_QNaN))) then (let done1 = True in - (FPProcessNaN type1 op1 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'N bits) . - (let (result :: 'N bits) = w__2 in + (FPProcessNaN type1 op1 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__2 :: ( 'N::len)Word.word) . + (let result = w__2 in return (done1, result)))) else if (((type2 = FPType_QNaN))) then (let done1 = True in - (FPProcessNaN type2 op2 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__3 :: 'N bits) . - (let (result :: 'N bits) = w__3 in + (FPProcessNaN type2 op2 fpcr :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'N::len)Word.word) . + (let result = w__3 in return (done1, result)))) else (let (done1 :: bool) = False in - (let (result :: 'N bits) = ((Zeros__1 ((int (size op1))) () :: ( 'N::len)Word.word)) in + (let result = ((Zeros__1 ((int (size op1))) () :: ( 'N::len)Word.word)) in return (done1, result))))))" @@ -8946,43 +8952,43 @@ definition HaveAArch32EL :: "(2)Word.word \<Rightarrow> bool " where definition AArch64_ResetSpecialRegisters :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " AArch64_ResetSpecialRegisters _ = ( - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . (write_reg SP_EL0_ref w__0 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (write_reg SP_EL1_ref w__1 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__2 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__2 :: 32 Word.word) . (write_reg SPSR_EL1_ref w__2 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__3 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__3 :: 64 Word.word) . ((((write_reg ELR_EL1_ref w__3 \<then> (if ((HaveEL EL2)) then - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 Word.word) . (write_reg SP_EL2_ref w__4 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__5 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__5 :: 32 Word.word) . (write_reg SPSR_EL2_ref w__5 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__6 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__6 :: 64 Word.word) . write_reg ELR_EL2_ref w__6))) else return () )) \<then> (if ((HaveEL EL3)) then - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 Word.word) . (write_reg SP_EL3_ref w__7 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__8 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__8 :: 32 Word.word) . (write_reg SPSR_EL3_ref w__8 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__9 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__9 :: 64 Word.word) . write_reg ELR_EL3_ref w__9))) else return () )) \<then> (if ((HaveAArch32EL EL1)) then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . (write_reg SPSR_fiq_ref w__10 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__11 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__11 :: 32 Word.word) . (write_reg SPSR_irq_ref w__11 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__12 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__12 :: 32 Word.word) . (write_reg SPSR_abt_ref w__12 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__13 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__13 :: 32 Word.word) . write_reg SPSR_und_ref w__13)))) else return () )) \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__14 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__14 :: 64 Word.word) . (write_reg DLR_EL0_ref w__14 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__15 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__15 :: 32 Word.word) . write_reg DSPSR_EL0_ref w__15)))))))" @@ -9013,6 +9019,7 @@ definition FPUnpackBase :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Rig (undefined_bitvector (( 5 :: int)::ii) :: ( 5 Word.word) M) \<bind> (\<lambda> (exp16 :: 5 bits) . (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (sign :: 1 bits) . (if (((((int (size fpval))) = (( 16 :: int)::ii)))) then + (let (fpval :: 16 Word.word) = ((Word.ucast fpval :: 16 Word.word)) in (let (sign :: 1 bits) = ((vec_of_bits [access_vec_dec fpval (( 15 :: int)::ii)] :: 1 Word.word)) in (let (exp16 :: 5 bits) = ((slice fpval (( 10 :: int)::ii) (( 5 :: int)::ii) :: 5 Word.word)) in (let (frac16 :: 10 bits) = ((slice fpval (( 0 :: int)::ii) (( 10 :: int)::ii) :: 10 Word.word)) in @@ -9057,8 +9064,9 @@ definition FPUnpackBase :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Rig ((realPowInteger (realFromFrac(( 20 :: int))(( 10 :: int))) ((- (( 10 :: int)::ii)))))))))) in (typ1, value_name)))) in (typ1, value_name))) in - return (sign, typ1, value_name))))) + return (sign, typ1, value_name)))))) else if (((((int (size fpval))) = (( 32 :: int)::ii)))) then + (let (fpval :: 32 Word.word) = ((Word.ucast fpval :: 32 Word.word)) in (let sign = ((vec_of_bits [access_vec_dec fpval (( 31 :: int)::ii)] :: 1 Word.word)) in (let exp32 = ((slice fpval (( 23 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) in (let frac32 = ((slice fpval (( 0 :: int)::ii) (( 23 :: int)::ii) :: 23 Word.word)) in @@ -9102,7 +9110,7 @@ definition FPUnpackBase :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Rig ((realPowInteger (realFromFrac(( 20 :: int))(( 10 :: int))) ((- (( 23 :: int)::ii)))))))))) in (typ1, value_name)))) in return (typ1, value_name))) \<bind> (\<lambda> varstup . (let ((typ1 :: FPType), (value_name :: real)) = varstup in - return (sign, typ1, value_name)))))) + return (sign, typ1, value_name))))))) else (let sign = ((vec_of_bits [access_vec_dec fpval (( 63 :: int)::ii)] :: 1 Word.word)) in (let exp64 = ((slice fpval (( 52 :: int)::ii) (( 11 :: int)::ii) :: 11 Word.word)) in @@ -9376,7 +9384,7 @@ definition FPToFixedJS :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow>(3 read_reg PSTATE_ref) \<bind> (\<lambda> (w__3 :: ProcState) . write_reg PSTATE_ref (w__3 (| ProcState_V := tup__3 |)))))))) else - (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . + (read_reg FPSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . write_reg FPSCR_ref ((set_slice (( 32 :: int)::ii) (( 4 :: int)::ii) w__4 (( 28 :: int)::ii) @@ -9385,7 +9393,7 @@ definition FPToFixedJS :: " int \<Rightarrow>('M::len)Word.word \<Rightarrow>(3 :: 4 Word.word)) :: 32 Word.word)))) \<then> return ((Word.ucast - ((GetSlice_int ((make_the_value (( 32 :: int)::ii) :: ( 'N::len)itself)) result (( 0 :: int)::ii) :: ( 'N::len)Word.word)) + ((GetSlice_int ((make_the_value (( 32 :: int)::ii) :: 32 itself)) result (( 0 :: int)::ii) :: 32 Word.word)) :: ( 'N::len)Word.word))))))))))))))))))))))" @@ -9425,9 +9433,8 @@ definition FPToFixed :: " int \<Rightarrow>('N::len)Word.word \<Rightarrow> int )) in (let (int_result :: ii) = (if round_up then ((ex_int int_result)) + (( 1 :: int)::ii) else int_result) in undefined_bool () \<bind> (\<lambda> (overflow :: bool) . - (undefined_bitvector M__tv :: (( 'M::len)Word.word) M) \<bind> (\<lambda> (result :: 'M bits) . - (SatQ int_result ((make_the_value ((int (size result))) :: ( 'M::len)itself)) unsigned - :: ((( 'M::len)Word.word * bool)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (undefined_bitvector M__tv :: (( 'M::len)Word.word) M) \<bind> (\<lambda> result . + (SatQ int_result ((make_the_value ((int (size result))) )) unsigned ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in (let result = tup__0 in (let overflow = tup__1 in (if overflow then FPProcessException FPExc_InvalidOp fpcr @@ -9448,7 +9455,7 @@ definition FPSqrt :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Rightarro (let typ1 = tup__0 in (let sign = tup__1 in (let value_name = tup__2 in - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . if ((((((typ1 = FPType_SNaN))) \<or> (((typ1 = FPType_QNaN)))))) then (FPProcessNaN typ1 op1 fpcr :: (( 'N::len)Word.word) M) else if (((typ1 = FPType_Zero))) then (FPZero ((int (size op1))) sign :: (( 'N::len)Word.word) M) @@ -9456,7 +9463,7 @@ definition FPSqrt :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Rightarro then (FPInfinity ((int (size op1))) sign :: (( 'N::len)Word.word) M) else if (((sign = (vec_of_bits [B1] :: 1 Word.word)))) then - (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__3 :: 'N bits) . + (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'N::len)Word.word) . (let result = w__3 in FPProcessException FPExc_InvalidOp fpcr \<then> return result)) else (FPRound__1 ((int (size op1))) ((sqrt value_name)) fpcr :: (( 'N::len)Word.word) M)))))))))))" @@ -9479,7 +9486,7 @@ definition FPRoundInt :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Right undefined_bool () \<bind> (\<lambda> (round_up :: bool) . undefined_real () \<bind> (\<lambda> (error :: real) . undefined_int () \<bind> (\<lambda> (int_result :: ii) . - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . if ((((((typ1 = FPType_SNaN))) \<or> (((typ1 = FPType_QNaN)))))) then (FPProcessNaN typ1 op1 fpcr :: (( 'N::len)Word.word) M) else if (((typ1 = FPType_Infinity))) then (FPInfinity ((int (size op1))) sign :: (( 'N::len)Word.word) M) @@ -9504,8 +9511,7 @@ definition FPRoundInt :: "('N::len)Word.word \<Rightarrow>(32)Word.word \<Right else int_result) in (let real_result = ((real_of_int int_result)) in (if (((real_result = (realFromFrac(( 0 :: int))(( 10 :: int)))))) then (FPZero ((int (size op1))) sign :: (( 'N::len)Word.word) M) - else (FPRound__0 ((int (size op1))) real_result fpcr FPRounding_ZERO :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> (result :: 'N - bits) . + else (FPRound__0 ((int (size op1))) real_result fpcr FPRounding_ZERO :: (( 'N::len)Word.word) M)) \<bind> (\<lambda> result . (if ((((((error \<noteq> (realFromFrac(( 0 :: int))(( 10 :: int)))))) \<and> exact))) then FPProcessException FPExc_Inexact fpcr else return () ) \<then> @@ -9566,9 +9572,9 @@ definition FPSub :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let type2 = tup__0 in (let sign2 = tup__1 in (let value2_name = tup__2 in - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . - (FPProcessNaNs type1 type2 op1 op2 fpcr :: ((bool * ( 'N::len)Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (FPProcessNaNs type1 type2 op1 op2 fpcr ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in (let done1 = tup__0 in (let result = tup__1 in (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (result_sign :: 1 bits) . @@ -9583,7 +9589,7 @@ definition FPSub :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let zero1 = (type1 = FPType_Zero) in (let zero2 = (type2 = FPType_Zero) in if ((((((inf1 \<and> inf2))) \<and> (((sign1 = sign2)))))) then - (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'N bits) . + (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'N::len)Word.word) . (let result = w__0 in FPProcessException FPExc_InvalidOp fpcr \<then> return result)) else if ((((((inf1 \<and> (((sign1 = (vec_of_bits [B0] :: 1 Word.word))))))) \<or> (((inf2 \<and> (((sign2 = (vec_of_bits [B1] :: 1 Word.word)))))))))) then @@ -9635,16 +9641,16 @@ definition FPMulAdd :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Ri (let (zero1 :: bool) = (type1 = FPType_Zero) in (let (inf2 :: bool) = (type2 = FPType_Infinity) in (let (zero2 :: bool) = (type2 = FPType_Zero) in - (undefined_bitvector ((int (size addend))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size addend))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . - (FPProcessNaNs3 typeA type1 type2 addend op1 op2 fpcr :: ((bool * ( 'N::len)Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (FPProcessNaNs3 typeA type1 type2 addend op1 op2 fpcr ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in (let done1 = tup__0 in (let result = tup__1 in (if ((((((typeA = FPType_QNaN))) \<and> ((((((inf1 \<and> zero2))) \<or> (((zero1 \<and> inf2))))))))) then - (FPDefaultNaN ((int (size addend))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'N bits) . + (FPDefaultNaN ((int (size addend))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'N::len)Word.word) . (let result = w__0 in FPProcessException FPExc_InvalidOp fpcr \<then> return result)) - else return result) \<bind> (\<lambda> (result :: 'N bits) . + else return result) \<bind> (\<lambda> result . (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (result_sign :: 1 bits) . undefined_real () \<bind> (\<lambda> (result_value :: real) . undefined_bool () \<bind> (\<lambda> (zeroP :: bool) . @@ -9659,7 +9665,7 @@ definition FPMulAdd :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Ri (let infP = (inf1 \<or> inf2) in (let zeroP = (zero1 \<or> zero2) in if (((((((((inf1 \<and> zero2))) \<or> (((zero1 \<and> inf2)))))) \<or> ((((((infA \<and> infP))) \<and> (((signA \<noteq> signP))))))))) then - (FPDefaultNaN ((int (size addend))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: 'N bits) . + (FPDefaultNaN ((int (size addend))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: ( 'N::len)Word.word) . (let result = w__1 in FPProcessException FPExc_InvalidOp fpcr \<then> return result)) else if ((((((infA \<and> (((signA = (vec_of_bits [B0] :: 1 Word.word))))))) \<or> (((infP \<and> (((signP = (vec_of_bits [B0] :: 1 Word.word)))))))))) then @@ -9698,9 +9704,9 @@ definition FPMul :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let type2 = tup__0 in (let sign2 = tup__1 in (let value2_name = tup__2 in - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . - (FPProcessNaNs type1 type2 op1 op2 fpcr :: ((bool * ( 'N::len)Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (FPProcessNaNs type1 type2 op1 op2 fpcr ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in (let done1 = tup__0 in (let result = tup__1 in undefined_bool () \<bind> (\<lambda> (zero2 :: bool) . @@ -9713,7 +9719,7 @@ definition FPMul :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let zero1 = (type1 = FPType_Zero) in (let zero2 = (type2 = FPType_Zero) in if ((((((inf1 \<and> zero2))) \<or> (((zero1 \<and> inf2)))))) then - (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'N bits) . + (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'N::len)Word.word) . (let result = w__0 in FPProcessException FPExc_InvalidOp fpcr \<then> return result)) else if (((inf1 \<or> inf2))) then @@ -9743,9 +9749,9 @@ definition FPMin :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let type2 = tup__0 in (let sign2 = tup__1 in (let value2_name = tup__2 in - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . - (FPProcessNaNs type1 type2 op1 op2 fpcr :: ((bool * ( 'N::len)Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (FPProcessNaNs type1 type2 op1 op2 fpcr ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in (let done1 = tup__0 in (let result = tup__1 in undefined_real () \<bind> (\<lambda> (value_name :: real) . @@ -9826,9 +9832,9 @@ definition FPMax :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let type2 = tup__0 in (let sign2 = tup__1 in (let value2_name = tup__2 in - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . - (FPProcessNaNs type1 type2 op1 op2 fpcr :: ((bool * ( 'N::len)Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (FPProcessNaNs type1 type2 op1 op2 fpcr ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in (let done1 = tup__0 in (let result = tup__1 in undefined_real () \<bind> (\<lambda> (value_name :: real) . @@ -9909,9 +9915,9 @@ definition FPDiv :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let type2 = tup__0 in (let sign2 = tup__1 in (let value2_name = tup__2 in - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . - (FPProcessNaNs type1 type2 op1 op2 fpcr :: ((bool * ( 'N::len)Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (FPProcessNaNs type1 type2 op1 op2 fpcr ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in (let done1 = tup__0 in (let result = tup__1 in undefined_bool () \<bind> (\<lambda> (zero2 :: bool) . @@ -9924,12 +9930,12 @@ definition FPDiv :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let zero1 = (type1 = FPType_Zero) in (let zero2 = (type2 = FPType_Zero) in if ((((((inf1 \<and> inf2))) \<or> (((zero1 \<and> zero2)))))) then - (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'N bits) . + (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'N::len)Word.word) . (let result = w__0 in FPProcessException FPExc_InvalidOp fpcr \<then> return result)) else if (((inf1 \<or> zero2))) then - (FPInfinity ((int (size op1))) ((xor_vec sign1 sign2 :: 1 Word.word)) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: 'N - bits) . + (FPInfinity ((int (size op1))) ((xor_vec sign1 sign2 :: 1 Word.word)) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__1 :: + ( 'N::len)Word.word) . (let result = w__1 in (if ((\<not> inf1)) then FPProcessException FPExc_DivideByZero fpcr else return () ) \<then> @@ -9960,9 +9966,9 @@ definition FPAdd :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let type2 = tup__0 in (let sign2 = tup__1 in (let value2_name = tup__2 in - (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (result :: 'N bits) . + (undefined_bitvector ((int (size op1))) :: (( 'N::len)Word.word) M) \<bind> (\<lambda> result . undefined_bool () \<bind> (\<lambda> (done1 :: bool) . - (FPProcessNaNs type1 type2 op1 op2 fpcr :: ((bool * ( 'N::len)Word.word)) M) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in + (FPProcessNaNs type1 type2 op1 op2 fpcr ) \<bind> (\<lambda> varstup . (let (tup__0, tup__1) = varstup in (let done1 = tup__0 in (let result = tup__1 in (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (result_sign :: 1 bits) . @@ -9977,7 +9983,7 @@ definition FPAdd :: "('N::len)Word.word \<Rightarrow>('N::len)Word.word \<Right (let zero1 = (type1 = FPType_Zero) in (let zero2 = (type2 = FPType_Zero) in if ((((((inf1 \<and> inf2))) \<and> (((sign1 = ((not_vec sign2 :: 1 Word.word)))))))) then - (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'N bits) . + (FPDefaultNaN ((int (size op1))) () :: (( 'N::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'N::len)Word.word) . (let result = w__0 in FPProcessException FPExc_InvalidOp fpcr \<then> return result)) else if ((((((inf1 \<and> (((sign1 = (vec_of_bits [B0] :: 1 Word.word))))))) \<or> (((inf2 \<and> (((sign2 = (vec_of_bits [B0] :: 1 Word.word)))))))))) then @@ -10122,23 +10128,23 @@ definition UpdateEDSCRFields :: " unit \<Rightarrow>((register_value),(unit),(e " UpdateEDSCRFields _ = ( Halted () \<bind> (\<lambda> (w__0 :: bool) . if ((\<not> w__0)) then - (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . (write_reg EDSCR_ref ((set_slice (( 32 :: int)::ii) (( 2 :: int)::ii) w__1 (( 8 :: int)::ii) (vec_of_bits [B0,B0] :: 2 Word.word) :: 32 Word.word)) \<then> - (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__2 :: 32 bits) . (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__3 :: 1 Word.word) . (write_reg EDSCR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__2 (( 18 :: int)::ii) w__3 :: 32 Word.word)) \<then> - (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__4 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__4 :: 32 bits) . write_reg EDSCR_ref ((set_slice (( 32 :: int)::ii) (( 4 :: int)::ii) w__4 (( 10 :: int)::ii) (vec_of_bits [B1,B1,B1,B1] :: 4 Word.word) :: 32 Word.word)))))) else - (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 bits) . read_reg PSTATE_ref \<bind> (\<lambda> (w__6 :: ProcState) . (write_reg EDSCR_ref ((set_slice (( 32 :: int)::ii) (( 2 :: int)::ii) w__5 (( 8 :: int)::ii)(ProcState_EL w__6) :: 32 Word.word)) \<then> - (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__7 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__7 :: 32 bits) . IsSecure () \<bind> (\<lambda> (w__8 :: bool) . (write_reg EDSCR_ref @@ -10210,7 +10216,7 @@ definition UpdateEDSCRFields :: " unit \<Rightarrow>((register_value),(unit),(e (let (RW :: 4 bits) = ((set_slice (( 4 :: int)::ii) (( 1 :: int)::ii) RW (( 0 :: int)::ii) w__19 :: 4 Word.word)) in return RW)) else return RW) \<bind> (\<lambda> (RW :: 4 bits) . - (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 bits) . write_reg EDSCR_ref ((set_slice (( 32 :: int)::ii) (( 4 :: int)::ii) w__20 (( 10 :: int)::ii) RW :: 32 Word.word))))))))))))))))))" @@ -10221,42 +10227,42 @@ definition Halt :: "(6)Word.word \<Rightarrow>((register_value),(unit),(excepti (CTI_SignalEvent CrossTriggerIn_CrossHalt \<then> UsingAArch32 () ) \<bind> (\<lambda> (w__0 :: bool) . ((if w__0 then - (ThisInstrAddr (( 32 :: int)::ii) () :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . + (ThisInstrAddr (( 32 :: int)::ii) () :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . (write_reg DLR_ref w__1 \<then> - (GetPSRFromPSTATE () :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__2 :: 32 bits) . + (GetPSRFromPSTATE () :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__2 :: 32 Word.word) . (write_reg DSPSR_ref w__2 \<then> - (read_reg DSPSR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__3 :: 32 Word.word) . + (read_reg DSPSR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__3 :: 32 bits) . read_reg PSTATE_ref \<bind> (\<lambda> (w__4 :: ProcState) . write_reg DSPSR_ref ((update_subrange_vec_dec w__3 (( 21 :: int)::ii) (( 21 :: int)::ii)(ProcState_SS w__4) :: 32 Word.word)))))) else - (ThisInstrAddr (( 64 :: int)::ii) () :: ( 64 Word.word) M) \<bind> (\<lambda> (w__5 :: 64 bits) . + (ThisInstrAddr (( 64 :: int)::ii) () :: ( 64 Word.word) M) \<bind> (\<lambda> (w__5 :: 64 Word.word) . (write_reg DLR_EL0_ref w__5 \<then> - (GetPSRFromPSTATE () :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__6 :: 32 bits) . + (GetPSRFromPSTATE () :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__6 :: 32 Word.word) . (write_reg DSPSR_EL0_ref w__6 \<then> - (read_reg DSPSR_EL0_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__7 :: 32 Word.word) . + (read_reg DSPSR_EL0_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__7 :: 32 bits) . read_reg PSTATE_ref \<bind> (\<lambda> (w__8 :: ProcState) . write_reg DSPSR_EL0_ref ((update_subrange_vec_dec w__7 (( 21 :: int)::ii) (( 21 :: int)::ii)(ProcState_SS w__8) :: 32 Word.word))))))) \<then> - (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__9 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__9 :: 32 bits) . (write_reg EDSCR_ref ((update_subrange_vec_dec w__9 (( 24 :: int)::ii) (( 24 :: int)::ii) (vec_of_bits [B1] :: 1 Word.word) :: 32 Word.word)) \<then> - (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__10 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__10 :: 32 bits) . (write_reg EDSCR_ref ((update_subrange_vec_dec w__10 (( 28 :: int)::ii) (( 28 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 32 Word.word)) \<then> IsSecure () ) \<bind> (\<lambda> (w__11 :: bool) . ((if w__11 then - (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 bits) . write_reg EDSCR_ref ((update_subrange_vec_dec w__12 (( 16 :: int)::ii) (( 16 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 32 Word.word))) else if ((HaveEL EL3)) then - (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__13 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__13 :: 32 bits) . ExternalSecureInvasiveDebugEnabled () \<bind> (\<lambda> (w__14 :: bool) . write_reg EDSCR_ref @@ -10266,14 +10272,14 @@ definition Halt :: "(6)Word.word \<Rightarrow>((register_value),(unit),(excepti :: 32 Word.word)))) else (read_reg EDSCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__15 :: 32 bits) . - assert_exp ((((vec_of_bits [access_vec_dec w__15 (( 16 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) ([(CHR ''(''), (CHR ''(''), (CHR ''E''), (CHR ''D''), (CHR ''S''), (CHR ''C''), (CHR ''R''), (CHR '')''), (CHR ''.''), (CHR ''S''), (CHR ''D''), (CHR ''D''), (CHR '' ''), (CHR ''=''), (CHR ''=''), (CHR '' ''), (char_of_nat 39), (CHR ''1''), (char_of_nat 39), (CHR '')'')]))) \<then> - (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__16 :: 32 Word.word) . + assert_exp ((((vec_of_bits [access_vec_dec w__15 (( 16 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))) ([(CHR ''(''), (CHR ''(''), (CHR ''E''), (CHR ''D''), (CHR ''S''), (CHR ''C''), (CHR ''R''), (CHR '')''), (CHR ''.''), (CHR ''S''), (CHR ''D''), (CHR ''D''), (CHR '' ''), (CHR ''=''), (CHR ''=''), (CHR '' ''), (CHR 0x27), (CHR ''1''), (CHR 0x27), (CHR '')'')]))) \<then> + (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__16 :: 32 bits) . (write_reg EDSCR_ref ((update_subrange_vec_dec w__16 (( 20 :: int)::ii) (( 20 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 32 Word.word)) \<then> UsingAArch32 () ) \<bind> (\<lambda> (w__17 :: bool) . ((if w__17 then - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__18 :: 4 bits) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__18 :: 4 Word.word) . (let split_vec = w__18 in (let (tup__0, tup__1, tup__2, tup__3) = ((subrange_vec_dec split_vec (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word), @@ -10295,7 +10301,7 @@ definition Halt :: "(6)Word.word \<Rightarrow>((register_value),(unit),(excepti read_reg PSTATE_ref) \<bind> (\<lambda> (w__24 :: ProcState) . write_reg PSTATE_ref (w__24 (| ProcState_T := ((vec_of_bits [B1] :: 1 Word.word))|))))))))))) else - (undefined_bitvector (( 5 :: int)::ii) :: ( 5 Word.word) M) \<bind> (\<lambda> (w__25 :: 5 bits) . + (undefined_bitvector (( 5 :: int)::ii) :: ( 5 Word.word) M) \<bind> (\<lambda> (w__25 :: 5 Word.word) . (let split_vec = w__25 in (let (tup__0, tup__1, tup__2, tup__3, tup__4) = ((subrange_vec_dec split_vec (( 4 :: int)::ii) (( 4 :: int)::ii) :: 1 Word.word), @@ -10316,7 +10322,7 @@ definition Halt :: "(6)Word.word \<Rightarrow>((register_value),(unit),(excepti read_reg PSTATE_ref) \<bind> (\<lambda> (w__31 :: ProcState) . ((write_reg PSTATE_ref (w__31 (| ProcState_IL := ((vec_of_bits [B0] :: 1 Word.word))|)) \<then> StopInstructionPrefetchAndEnableITR () ) \<then> - (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__32 :: 32 Word.word) . + (read_reg EDSCR_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__32 :: 32 bits) . write_reg EDSCR_ref ((update_subrange_vec_dec w__32 (( 5 :: int)::ii) (( 0 :: int)::ii) reason :: 32 Word.word)) \<then> UpdateEDSCRFields () )))))))))" @@ -10709,7 +10715,7 @@ definition AArch64_S1AttrDecode :: "(2)Word.word \<Rightarrow>(3)Word.word \<Ri (vec_of_bits [B0,B0,B1,B1] :: 4 Word.word) :: 4 Word.word)) \<noteq> (vec_of_bits [B0,B0,B0,B0] :: 4 Word.word)))))))))) then (ConstrainUnpredictableBits (( 8 :: int)::ii) Unpredictable_RESMAIR :: ((Constraint * 8 Word.word)) M) \<bind> (\<lambda> (w__0 :: - (Constraint * 8 bits)) . + (Constraint * 8 Word.word)) . (let (tup__0, tup__1) = w__0 in (let (anon10 :: Constraint) = tup__0 in return tup__1))) @@ -10810,7 +10816,7 @@ definition AArch64_CheckAndUpdateDescriptor_SecondStage :: " DescriptorUpdate \ (let descaddr2 = ((DescriptorUpdate_descaddr result)) in CreateAccessDescriptor AccType_ATOMICRW \<bind> (\<lambda> (w__0 :: AccessDescriptor) . (let accdesc = w__0 in - (aget__Mem descaddr2 (( 8 :: int)::ii) accdesc :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . + (aget__Mem descaddr2 (( 8 :: int)::ii) accdesc :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (let desc = w__1 in (let (desc :: 64 bits) = (if hw_update_AF then @@ -10891,7 +10897,7 @@ definition AArch64_CheckS2Permission :: " Permissions \<Rightarrow>(64)Word.wor undefined_bool () \<bind> (\<lambda> (secondstage :: bool) . (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (domain1 :: 4 bits) . if fail1 then - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__8 :: 4 bits) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__8 :: 4 Word.word) . (let domain1 = w__8 in (let secondstage = True in AArch64_PermissionFault ipaddress level acctype ((\<not> failedread)) secondstage s2fs1walk))) @@ -10912,9 +10918,9 @@ definition ZeroExtend_slice_append :: " int \<Rightarrow>('n::len)Word.word \<R " ZeroExtend_slice_append (o__tv :: int) xs i l ys = ( assert_exp True ('''') \<then> ((let xs = ((and_vec xs ((slice_mask ((int (size xs))) i l :: ( 'n::len)Word.word)) :: ( 'n::len)Word.word)) in - (let (xs :: 'o bits) = + (let xs = ((shiftl ((extzv o__tv ((shiftr xs i :: ( 'n::len)Word.word)) :: ( 'o::len)Word.word)) ((int (size ys))) :: ( 'o::len)Word.word)) in - (let (ys :: 'o bits) = ((extzv ((int (size xs))) ys :: ( 'o::len)Word.word)) in + (let ys = ((extzv ((int (size xs))) ys :: ( 'o::len)Word.word)) in return ((or_vec xs ys :: ( 'o::len)Word.word)))))))" @@ -10958,7 +10964,7 @@ definition AArch64_TranslationTableWalk_SecondStage :: "(52)Word.word \<Rightar liftR (undefined_bool () ) \<bind> (\<lambda> (midgrain :: bool) . liftR (undefined_bool () ) \<bind> (\<lambda> (largegrain :: bool) . liftR (undefined_int () ) \<bind> (\<lambda> (top1 :: ii) . - liftR ((ZeroExtend__1 (( 64 :: int)::ii) ipaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__6 :: 64 bits) . + liftR ((ZeroExtend__1 (( 64 :: int)::ii) ipaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__6 :: 64 Word.word) . (let inputaddr = w__6 in liftR ((read_reg VTCR_EL2_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__7 :: 32 bits) . (let inputsize = ((( 64 :: int)::ii) - ((Word.uint ((slice w__7 (( 0 :: int)::ii) (( 6 :: int)::ii) :: 6 Word.word))))) in @@ -11209,13 +11215,13 @@ definition AArch64_TranslationTableWalk_SecondStage :: "(52)Word.word \<Rightar (let tmp_220 = ((tmp_220 (| FullAddress_NS := ns_table |))) in (let descaddr = ((descaddr (| AddressDescriptor_paddress := tmp_220 |))) in (let descaddr2 = descaddr in - liftR ((ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__34 :: 64 bits) . + liftR ((ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__34 :: 64 Word.word) . (let descaddr2 = ((descaddr2 (| AddressDescriptor_vaddress := w__34 |))) in liftR (CreateAccessDescriptorPTW acctype True s2fs1walk level) \<bind> (\<lambda> (w__35 :: AccessDescriptor) . (let accdesc = w__35 in - liftR ((aget__Mem descaddr2 (( 8 :: int)::ii) accdesc :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__36 :: 64 - bits) . + liftR ((aget__Mem descaddr2 (( 8 :: int)::ii) accdesc :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__36 :: + 64 Word.word) . (let desc = w__36 in (if reversedescriptors then liftR ((BigEndianReverse desc :: ( 64 Word.word) M)) else return desc) \<bind> (\<lambda> (desc :: 64 bits) . @@ -11395,7 +11401,7 @@ definition AArch64_TranslationTableWalk_SecondStage :: "(52)Word.word \<Rightar (vec_of_bits [B1] :: 1 Word.word) :: 3 Word.word)) in (let (memattr :: 4 bits) = ((slice desc (( 2 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word)) in - liftR ((undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M)) \<bind> (\<lambda> (w__51 :: 4 bits) . + liftR ((undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M)) \<bind> (\<lambda> (w__51 :: 4 Word.word) . (let result = ((result (| TLBRecord_domain := w__51 |))) in (let result = ((result (| TLBRecord_level := level |))) in (let result = @@ -12014,22 +12020,21 @@ definition aarch64_integer_conditional_compare_immediate :: "(4)Word.word \<Rig (assert_exp True (''datasize constraint'') \<then> assert_exp True (''dbytes constraint'')) \<then> ((let flags = flags__arg in - (aget_X datasize n :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (operand1 :: 'datasize bits) . - (let (operand2 :: 'datasize bits) = imm in + (aget_X datasize n :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> operand1 . + (let operand2 = imm in (let (carry_in :: 1 bits) = ((vec_of_bits [B0] :: 1 Word.word)) in - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (anon10 :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> anon10 . ConditionHolds condition \<bind> (\<lambda> (w__0 :: bool) . (let (flags :: 4 Word.word) = (if w__0 then - (let ((carry_in :: 1 bits), (operand2 :: 'datasize bits)) = + (let ((carry_in :: 1 bits), operand2) = (if sub_op then - (let (operand2 :: 'datasize bits) = ((not_vec operand2 :: ( 'datasize::len)Word.word)) in + (let operand2 = ((not_vec operand2 :: ( 'datasize::len)Word.word)) in (let (carry_in :: 1 bits) = ((vec_of_bits [B1] :: 1 Word.word)) in (carry_in, operand2))) else (carry_in, operand2)) in - (let (tup__0, tup__1) = - ((AddWithCarry operand1 operand2 carry_in :: (( 'datasize::len)Word.word * 4 Word.word))) in - (let (anon10 :: 'datasize bits) = tup__0 in + (let (tup__0, tup__1) = (AddWithCarry operand1 operand2 carry_in ) in + (let anon10 = tup__0 in tup__1))) else flags) in (let (tup__0, tup__1, tup__2, tup__3) = @@ -12081,7 +12086,7 @@ definition ConditionSyndrome :: " unit \<Rightarrow>((register_value),((5)Word. (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (cond :: 4 bits) . UsingAArch32 () \<bind> (\<lambda> (w__0 :: bool) . if w__0 then - (AArch32_CurrentCond () :: ( 4 Word.word) M) \<bind> (\<lambda> (w__1 :: 4 bits) . + (AArch32_CurrentCond () :: ( 4 Word.word) M) \<bind> (\<lambda> (w__1 :: 4 Word.word) . (let cond = w__1 in read_reg PSTATE_ref \<bind> (\<lambda> (w__2 :: ProcState) . if ((((ProcState_T w__2) = (vec_of_bits [B0] :: 1 Word.word)))) then @@ -12122,10 +12127,11 @@ definition BranchToAddr :: "('N::len)Word.word \<Rightarrow> BranchType \<Right write_reg BranchTaken_ref True \<then> ((let (_ :: unit) = (Hint_Branch branch_type) in if (((((int (size target))) = (( 32 :: int)::ii)))) then + (let (target :: 32 Word.word) = ((Word.ucast target :: 32 Word.word)) in UsingAArch32 () \<bind> (\<lambda> (w__0 :: bool) . (assert_exp w__0 (''UsingAArch32()'') \<then> - (ZeroExtend__1 (( 64 :: int)::ii) target :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 bits) . - write_reg PC_ref w__1)) + (ZeroExtend__1 (( 64 :: int)::ii) target :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + write_reg PC_ref w__1))) else and_boolM (return (((((int (size target))) = (( 64 :: int)::ii))))) (UsingAArch32 () \<bind> (\<lambda> (w__2 :: bool) . return ((\<not> w__2)))) \<bind> (\<lambda> (w__3 :: bool) . @@ -12170,13 +12176,13 @@ definition aset_Rmode :: " int \<Rightarrow>(5)Word.word \<Rightarrow>(32)Word. read_reg R_ref \<bind> (\<lambda> (w__2 :: ( 64 bits) list) . (let (tmp_10 :: 64 bits) = ((access_list_dec w__2 n :: 64 Word.word)) in (let tmp_10 = ((update_subrange_vec_dec tmp_10 (( 31 :: int)::ii) (( 0 :: int)::ii) value_name :: 64 Word.word)) in - read_reg R_ref \<bind> (\<lambda> (w__3 :: ( 64 Word.word) list) . + read_reg R_ref \<bind> (\<lambda> (w__3 :: ( 64 bits) list) . write_reg R_ref ((update_list_dec w__3 n tmp_10 :: ( 64 Word.word) list)))))) else and_boolM (return ((\<not> ((HighestELUsingAArch32 () ))))) ((ConstrainUnpredictableBool Unpredictable_ZEROUPPER)) \<bind> (\<lambda> (w__5 :: bool) . if w__5 then - read_reg R_ref \<bind> (\<lambda> (w__6 :: ( 64 Word.word) list) . + read_reg R_ref \<bind> (\<lambda> (w__6 :: ( 64 bits) list) . LookUpRIndex n mode \<bind> (\<lambda> (w__7 :: ii) . (ZeroExtend__0 value_name ((make_the_value (( 64 :: int)::ii) :: 64 itself)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__8 :: 64 Word.word) . @@ -12186,7 +12192,7 @@ definition aset_Rmode :: " int \<Rightarrow>(5)Word.word \<Rightarrow>(32)Word. LookUpRIndex n mode \<bind> (\<lambda> (w__10 :: ii) . (let (tmp_20 :: 64 bits) = ((access_list_dec w__9 w__10 :: 64 Word.word)) in (let tmp_20 = ((update_subrange_vec_dec tmp_20 (( 31 :: int)::ii) (( 0 :: int)::ii) value_name :: 64 Word.word)) in - read_reg R_ref \<bind> (\<lambda> (w__11 :: ( 64 Word.word) list) . + read_reg R_ref \<bind> (\<lambda> (w__11 :: ( 64 bits) list) . LookUpRIndex n mode \<bind> (\<lambda> (w__12 :: ii) . write_reg R_ref ((update_list_dec w__11 w__12 tmp_20 :: ( 64 Word.word) list)))))))))))))" @@ -12548,7 +12554,7 @@ definition AddPAC :: "(64)Word.word \<Rightarrow>(64)Word.word \<Rightarrow>(12 :: 64 Word.word)) in (ComputePAC ext_ptr modifier ((subrange_vec_dec K (( 127 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) ((subrange_vec_dec K (( 63 :: int)::ii) (( 0 :: int)::ii) :: 64 Word.word)) - :: ( 64 Word.word) M) \<bind> (\<lambda> (w__27 :: 64 bits) . + :: ( 64 Word.word) M) \<bind> (\<lambda> (w__27 :: 64 Word.word) . (let (PAC :: 64 bits) = w__27 in (let (PAC :: 64 bits) = (if (((((\<not> ((is_zero_subrange ptr @@ -12631,16 +12637,16 @@ definition AArch64_vESBOperation :: " unit \<Rightarrow>((register_value),(unit (read_reg VDFSR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 bits) . (AArch32_ReportDeferredSError ((slice w__13 (( 14 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) (vec_of_bits [access_vec_dec w__14 (( 12 :: int)::ii)] :: 1 Word.word) - :: ( 32 Word.word) M) \<bind> (\<lambda> (w__15 :: 32 bits) . + :: ( 32 Word.word) M) \<bind> (\<lambda> (w__15 :: 32 Word.word) . (let (VDISR :: 32 bits) = w__15 in return () )))) else (read_reg VSESR_EL2_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: 32 bits) . - (AArch64_ReportDeferredSError ((slice w__16 (( 0 :: int)::ii) (( 25 :: int)::ii) :: 25 Word.word)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: 64 - bits) . + (AArch64_ReportDeferredSError ((slice w__16 (( 0 :: int)::ii) (( 25 :: int)::ii) :: 25 Word.word)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: + 64 Word.word) . (let (VDISR_EL2 :: 64 bits) = w__17 in return () )))) \<then> - (read_reg HCR_EL2_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__18 :: 64 Word.word) . + (read_reg HCR_EL2_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__18 :: 64 bits) . write_reg HCR_EL2_ref ((set_slice (( 64 :: int)::ii) (( 1 :: int)::ii) w__18 (( 8 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 64 Word.word)))) @@ -12757,7 +12763,7 @@ definition AArch64_WatchpointByteMatch :: " int \<Rightarrow>(64)Word.word \<Ri definition IsOnes_slice :: "('n::len)Word.word \<Rightarrow> int \<Rightarrow> int \<Rightarrow>((register_value),(bool),(exception))monad " where " IsOnes_slice xs i l = ( assert_exp True ('''') \<then> - ((let (m :: 'n bits) = ((slice_mask ((int (size xs))) i l :: ( 'n::len)Word.word)) in + ((let m = ((slice_mask ((int (size xs))) i l :: ( 'n::len)Word.word)) in return (((((and_vec xs m :: ( 'n::len)Word.word)) = m))))))" @@ -12811,7 +12817,7 @@ definition AArch64_TranslationTableWalk :: "(52)Word.word \<Rightarrow>(64)Word liftR (undefined_bool () ) \<bind> (\<lambda> (largegrain :: bool) . liftR (undefined_int () ) \<bind> (\<lambda> (top1 :: ii) . (if ((\<not> secondstage)) then - liftR ((ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__8 :: 64 bits) . + liftR ((ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__8 :: 64 Word.word) . (let inputaddr = w__8 in liftR (read_reg PSTATE_ref) \<bind> (\<lambda> (w__9 :: ProcState) . liftR (AddrTop inputaddr (((acctype = AccType_IFETCH)))(ProcState_EL w__9)) \<bind> (\<lambda> (w__10 :: @@ -13338,7 +13344,7 @@ definition AArch64_TranslationTableWalk :: "(52)Word.word \<Rightarrow>(64)Word update_AF, update_AP)))))))))))) else - liftR ((ZeroExtend__1 (( 64 :: int)::ii) ipaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__115 :: 64 bits) . + liftR ((ZeroExtend__1 (( 64 :: int)::ii) ipaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__115 :: 64 Word.word) . (let inputaddr = w__115 in liftR ((read_reg VTCR_EL2_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__116 :: 32 bits) . (let inputsize = ((( 64 :: int)::ii) - ((Word.uint ((slice w__116 (( 0 :: int)::ii) (( 6 :: int)::ii) :: 6 Word.word))))) in @@ -13647,13 +13653,14 @@ definition AArch64_TranslationTableWalk :: "(52)Word.word \<Rightarrow>(64)Word else return result) \<bind> (\<lambda> (result :: TLBRecord) . return (descaddr2, hwupdatewalk, result)))))) \<bind> (\<lambda> varstup . (let ((descaddr2 :: AddressDescriptor), (hwupdatewalk :: bool), (result :: TLBRecord)) = varstup in - liftR ((ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__146 :: 64 bits) . + liftR ((ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__146 :: + 64 Word.word) . (let descaddr2 = ((descaddr2 (| AddressDescriptor_vaddress := w__146 |))) in liftR (CreateAccessDescriptorPTW acctype secondstage s2fs1walk level) \<bind> (\<lambda> (w__147 :: AccessDescriptor) . (let accdesc = w__147 in - liftR ((aget__Mem descaddr2 (( 8 :: int)::ii) accdesc :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__148 :: 64 - bits) . + liftR ((aget__Mem descaddr2 (( 8 :: int)::ii) accdesc :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__148 :: + 64 Word.word) . (let desc = w__148 in (if reversedescriptors then liftR ((BigEndianReverse desc :: ( 64 Word.word) M)) else return desc) \<bind> (\<lambda> (desc :: 64 bits) . @@ -13934,7 +13941,7 @@ definition AArch64_TranslationTableWalk :: "(52)Word.word \<Rightarrow>(64)Word (vec_of_bits [B1] :: 1 Word.word) :: 3 Word.word)) in (let (memattr :: 4 bits) = ((slice desc (( 2 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word)) in - liftR ((undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M)) \<bind> (\<lambda> (w__163 :: 4 bits) . + liftR ((undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M)) \<bind> (\<lambda> (w__163 :: 4 Word.word) . (let result = ((result (| TLBRecord_domain := w__163 |))) in (let result = ((result (| TLBRecord_level := level |))) in (let result = @@ -14111,7 +14118,7 @@ definition AArch64_TranslateAddressS1Off :: "(64)Word.word \<Rightarrow> AccTyp bool) . if ((\<not> w__3)) then (let level = ((( 0 :: int)::ii)) in - (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__4 :: 52 bits) . + (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__4 :: 52 Word.word) . (let ipaddress = w__4 in (let secondstage = False in (let s2fs1walk = False in @@ -14271,7 +14278,7 @@ definition AArch64_TranslateAddressS1Off :: "(64)Word.word \<Rightarrow> AccTyp (let tmp_2370 = ((tmp_2370 (| AddressDescriptor_memattrs := w__10 |))) in (let result = ((result (| TLBRecord_addrdesc := tmp_2370 |))) in (let (tmp_2380 :: Permissions) = ((TLBRecord_perms result)) in - (undefined_bitvector (( 3 :: int)::ii) :: ( 3 Word.word) M) \<bind> (\<lambda> (w__11 :: 3 bits) . + (undefined_bitvector (( 3 :: int)::ii) :: ( 3 Word.word) M) \<bind> (\<lambda> (w__11 :: 3 Word.word) . (let tmp_2380 = ((tmp_2380 (| Permissions_ap := w__11 |))) in (let result = ((result (| TLBRecord_perms := tmp_2380 |))) in (let (tmp_2390 :: Permissions) = ((TLBRecord_perms result)) in @@ -14280,11 +14287,11 @@ definition AArch64_TranslateAddressS1Off :: "(64)Word.word \<Rightarrow> AccTyp (let (tmp_2400 :: Permissions) = ((TLBRecord_perms result)) in (let tmp_2400 = ((tmp_2400 (| Permissions_pxn := ((vec_of_bits [B0] :: 1 Word.word))|))) in (let result = ((result (| TLBRecord_perms := tmp_2400 |))) in - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__12 :: 1 bits) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__12 :: 1 Word.word) . (let result = ((result (| TLBRecord_nG := w__12 |))) in undefined_bool () \<bind> (\<lambda> (w__13 :: bool) . (let result = ((result (| TLBRecord_contiguous := w__13 |))) in - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__14 :: 4 bits) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__14 :: 4 Word.word) . (let result = ((result (| TLBRecord_domain := w__14 |))) in undefined_int () \<bind> (\<lambda> (w__15 :: ii) . (let result = ((result (| TLBRecord_level := w__15 |))) in @@ -14365,7 +14372,7 @@ definition AArch64_MaybeZeroRegisterUppers :: " unit \<Rightarrow>((register_va (let tmp_30 = ((set_slice (( 64 :: int)::ii) (( 32 :: int)::ii) tmp_30 (( 32 :: int)::ii) ((Zeros__1 (( 32 :: int)::ii) () :: 32 Word.word)) :: 64 Word.word)) in - read_reg R_ref \<bind> (\<lambda> (w__15 :: ( 64 Word.word) list) . + read_reg R_ref \<bind> (\<lambda> (w__15 :: ( 64 bits) list) . write_reg R_ref ((update_list_dec w__15 n tmp_30 :: ( 64 Word.word) list)))))) else return () )))))))))))" @@ -14420,7 +14427,7 @@ definition DCPSInstruction :: "(2)Word.word \<Rightarrow>((register_value),(uni (((if w__16 then read_reg PSTATE_ref \<bind> (\<lambda> (w__17 :: ProcState) . ((if ((((ProcState_M w__17) = M32_Monitor))) then - (read_reg SCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 Word.word) . + (read_reg SCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 bits) . write_reg SCR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__18 (( 0 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 32 Word.word))) @@ -14452,9 +14459,9 @@ definition DCPSInstruction :: "(2)Word.word \<Rightarrow>((register_value),(uni else return () ) else return () )) \<then> (if (((handle_el = EL2))) then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__26 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__26 :: 32 Word.word) . (write_reg ELR_hyp_ref w__26 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__27 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__27 :: 32 Word.word) . write_reg HSR_ref w__27)) else (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__28 :: 32 Word.word) . @@ -14466,9 +14473,9 @@ definition DCPSInstruction :: "(2)Word.word \<Rightarrow>((register_value),(uni (write_reg PSTATE_ref (w__30 (| ProcState_E := ((vec_of_bits [access_vec_dec w__31 (( 25 :: int)::ii)] :: 1 Word.word))|)) \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__32 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__32 :: 32 Word.word) . (write_reg DLR_ref w__32 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__33 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__33 :: 32 Word.word) . write_reg DSPSR_ref w__33))))))))) else UsingAArch32 () \<bind> (\<lambda> (w__34 :: bool) . @@ -14504,9 +14511,9 @@ definition DCPSInstruction :: "(2)Word.word \<Rightarrow>((register_value),(uni (aset_SPSR w__50 \<then> (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__51 :: 32 Word.word) . (aset_ESR__1 w__51 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__52 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__52 :: 64 Word.word) . (write_reg DLR_EL0_ref w__52 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__53 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__53 :: 32 Word.word) . write_reg DSPSR_EL0_ref w__53 \<then> (if ((HaveUAOExt () )) then read_reg PSTATE_ref \<bind> (\<lambda> (w__54 :: ProcState) . @@ -14623,9 +14630,9 @@ definition AArch64_AbortSyndrome :: " Exception \<Rightarrow> FaultRecord \<Rig " AArch64_AbortSyndrome typ1 fault vaddress = ( ExceptionSyndrome typ1 \<bind> (\<lambda> (exception :: ExceptionRecord) . (let (d_side :: bool) = ((((typ1 = Exception_DataAbort))) \<or> (((typ1 = Exception_Watchpoint)))) in - (AArch64_FaultSyndrome d_side fault :: ( 25 Word.word) M) \<bind> (\<lambda> (w__0 :: 25 bits) . + (AArch64_FaultSyndrome d_side fault :: ( 25 Word.word) M) \<bind> (\<lambda> (w__0 :: 25 Word.word) . (let exception = ((exception (| ExceptionRecord_syndrome := w__0 |))) in - (ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . + (ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (let exception = ((exception (| ExceptionRecord_vaddress := w__1 |))) in IPAValid fault \<bind> (\<lambda> (w__2 :: bool) . (let (exception :: ExceptionRecord) = @@ -14667,7 +14674,7 @@ definition AArch64_ExceptionClass :: " Exception \<Rightarrow>(2)Word.word \<Ri (if (((((ex_int w__0)) = (( 32 :: int)::ii)))) then (vec_of_bits [B1] :: 1 Word.word) else (vec_of_bits [B0] :: 1 Word.word)) in UsingAArch32 () \<bind> (\<lambda> (from_32 :: bool) . - (assert_exp (((from_32 \<or> (((il = (vec_of_bits [B1] :: 1 Word.word))))))) ([(CHR ''(''), (CHR ''f''), (CHR ''r''), (CHR ''o''), (CHR ''m''), (CHR ''_''), (CHR ''3''), (CHR ''2''), (CHR '' ''), (CHR ''|''), (CHR ''|''), (CHR '' ''), (CHR ''(''), (CHR ''i''), (CHR ''l''), (CHR '' ''), (CHR ''=''), (CHR ''=''), (CHR '' ''), (char_of_nat 39), (CHR ''1''), (char_of_nat 39), (CHR '')''), (CHR '')'')]) \<then> + (assert_exp (((from_32 \<or> (((il = (vec_of_bits [B1] :: 1 Word.word))))))) ([(CHR ''(''), (CHR ''f''), (CHR ''r''), (CHR ''o''), (CHR ''m''), (CHR ''_''), (CHR ''3''), (CHR ''2''), (CHR '' ''), (CHR ''|''), (CHR ''|''), (CHR '' ''), (CHR ''(''), (CHR ''i''), (CHR ''l''), (CHR '' ''), (CHR ''=''), (CHR ''=''), (CHR '' ''), (CHR 0x27), (CHR ''1''), (CHR 0x27), (CHR '')''), (CHR '')'')]) \<then> undefined_int () ) \<bind> (\<lambda> (ec :: ii) . (case typ1 of Exception_Uncategorized => @@ -14797,14 +14804,14 @@ definition AArch64_ReportException :: " ExceptionRecord \<Rightarrow>(2)Word.wo aset_FAR__0 target_el w__0))) \<then> (if (((target_el = EL2))) then if(ExceptionRecord_ipavalid exception) then - (read_reg HPFAR_EL2_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (read_reg HPFAR_EL2_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . write_reg HPFAR_EL2_ref ((set_slice (( 64 :: int)::ii) (( 40 :: int)::ii) w__1 (( 4 :: int)::ii) ((slice(ExceptionRecord_ipaddress exception) (( 12 :: int)::ii) (( 40 :: int)::ii) :: 40 Word.word)) :: 64 Word.word))) else - (read_reg HPFAR_EL2_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . + (read_reg HPFAR_EL2_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . (undefined_bitvector (( 40 :: int)::ii) :: ( 40 Word.word) M) \<bind> (\<lambda> (w__3 :: 40 Word.word) . write_reg HPFAR_EL2_ref ((set_slice (( 64 :: int)::ii) (( 40 :: int)::ii) w__2 (( 4 :: int)::ii) w__3 :: 64 Word.word)))) else return () )))))))))))" @@ -14868,14 +14875,14 @@ definition AArch64_ESBOperation :: " unit \<Rightarrow>((register_value),(unit) (let syndrome32 = w__25 in (AArch32_ReportDeferredSError(AArch32_SErrorSyndrome_AET syndrome32)(AArch32_SErrorSyndrome_ExT syndrome32) - :: ( 32 Word.word) M) \<bind> (\<lambda> (w__26 :: 32 bits) . + :: ( 32 Word.word) M) \<bind> (\<lambda> (w__26 :: 32 Word.word) . (let (DISR :: 32 bits) = w__26 in return () )))) else (let implicit_esb = False in - (AArch64_PhysicalSErrorSyndrome implicit_esb :: ( 25 Word.word) M) \<bind> (\<lambda> (w__27 :: 25 bits) . + (AArch64_PhysicalSErrorSyndrome implicit_esb :: ( 25 Word.word) M) \<bind> (\<lambda> (w__27 :: 25 Word.word) . (let syndrome64 = w__27 in - (AArch64_ReportDeferredSError syndrome64 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: 64 bits) . + (AArch64_ReportDeferredSError syndrome64 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: 64 Word.word) . (let (DISR_EL1 :: 64 bits) = w__28 in return () )))))) \<then> ClearPendingPhysicalSError () )) @@ -14924,7 +14931,7 @@ definition AArch64_CheckAndUpdateDescriptor :: " DescriptorUpdate \<Rightarrow> return descaddr2)))) \<bind> (\<lambda> (descaddr2 :: AddressDescriptor) . liftR (CreateAccessDescriptor AccType_ATOMICRW) \<bind> (\<lambda> (w__4 :: AccessDescriptor) . (let accdesc = w__4 in - liftR ((aget__Mem descaddr2 (( 8 :: int)::ii) accdesc :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__5 :: 64 bits) . + liftR ((aget__Mem descaddr2 (( 8 :: int)::ii) accdesc :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__5 :: 64 Word.word) . (let desc = w__5 in (if reversedescriptors then liftR ((BigEndianReverse desc :: ( 64 Word.word) M)) else return desc) \<bind> (\<lambda> (desc :: 64 bits) . @@ -14984,7 +14991,7 @@ definition AArch64_StateMatch :: "(2)Word.word \<Rightarrow>(1)Word.word \<Righ :: 5 Word.word)) = (vec_of_bits [B1,B1,B1,B1,B0] :: 5 Word.word))))))))))))))))))) \<or> (((((((((HMC = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((PxC = (vec_of_bits [B0,B0] :: 2 Word.word))))))) \<and> (((((\<not> isbreakpnt)) \<or> ((\<not> ((HaveAArch32EL EL1))))))))))))) \<or> (((((((((SSC = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<or> (((SSC = (vec_of_bits [B1,B0] :: 2 Word.word))))))) \<and> ((\<not> ((HaveEL EL3)))))))))) \<or> ((((((((((((((concat_vec HMC SSC :: 3 Word.word)) \<noteq> (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((concat_vec HMC SSC :: 3 Word.word)) \<noteq> (vec_of_bits [B1,B1,B1] :: 3 Word.word))))))) \<and> ((\<not> ((HaveEL EL3))))))) \<and> ((\<not> ((HaveEL EL2)))))))))) \<or> ((((((((concat_vec ((concat_vec HMC SSC :: 3 Word.word)) PxC :: 5 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B0] :: 5 Word.word)))) \<and> ((\<not> ((HaveEL EL2)))))))))) then liftR ((undefined_bitvector (( 5 :: int)::ii) :: ( 5 Word.word) M)) \<bind> (\<lambda> (tmp_50 :: 5 bits) . liftR ((ConstrainUnpredictableBits (( 5 :: int)::ii) Unpredictable_RESBPWPCTRL - :: ((Constraint * 5 Word.word)) M)) \<bind> (\<lambda> (w__0 :: (Constraint * 5 bits)) . + :: ((Constraint * 5 Word.word)) M)) \<bind> (\<lambda> (w__0 :: (Constraint * 5 Word.word)) . (let (tup__0, tup__1) = w__0 in (let c = tup__0 in (let tmp_50 = tup__1 in @@ -15055,7 +15062,7 @@ definition AArch64_StateMatch :: "(2)Word.word \<Rightarrow>(1)Word.word \<Righ liftR (undefined_bool () ) \<bind> (\<lambda> (linked_to :: bool) . liftR ((undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (vaddress :: 64 bits) . (if linked then - liftR ((undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__7 :: 64 bits) . + liftR ((undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__7 :: 64 Word.word) . (let (vaddress :: 64 bits) = w__7 in (let (linked_to :: bool) = True in (let (linked_match :: bool) = (AArch64_BreakpointValueMatch lbn vaddress linked_to) in @@ -15231,16 +15238,17 @@ definition BranchTo :: "('N::len)Word.word \<Rightarrow> BranchType \<Rightarro write_reg BranchTaken_ref True \<then> ((let (_ :: unit) = (Hint_Branch branch_type) in if (((((int (size target))) = (( 32 :: int)::ii)))) then + (let (target :: 32 Word.word) = ((Word.ucast target :: 32 Word.word)) in UsingAArch32 () \<bind> (\<lambda> (w__0 :: bool) . (assert_exp w__0 (''UsingAArch32()'') \<then> - (ZeroExtend__1 (( 64 :: int)::ii) target :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 bits) . - write_reg PC_ref w__1)) + (ZeroExtend__1 (( 64 :: int)::ii) target :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + write_reg PC_ref w__1))) else and_boolM (return (((((int (size target))) = (( 64 :: int)::ii))))) (UsingAArch32 () \<bind> (\<lambda> (w__2 :: bool) . return ((\<not> w__2)))) \<bind> (\<lambda> (w__3 :: bool) . (assert_exp w__3 (''((N == 64) && !(UsingAArch32()))'') \<then> - (AArch64_BranchAddr ((slice target (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__4 :: 64 - bits) . + (AArch64_BranchAddr ((slice target (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__4 :: + 64 Word.word) . write_reg PC_ref w__4)))))" @@ -16937,7 +16945,7 @@ definition AArch64_VectorCatchException :: " FaultRecord \<Rightarrow>((registe return ((((vec_of_bits [access_vec_dec w__3 (( 27 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))))) ((read_reg MDCR_EL2_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . return ((((vec_of_bits [access_vec_dec w__4 (( 8 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)))))))) \<bind> (\<lambda> (w__6 :: bool) . - (assert_exp w__6 ([(CHR ''(''), (CHR ''(''), (CHR ''H''), (CHR ''a''), (CHR ''v''), (CHR ''e''), (CHR ''E''), (CHR ''L''), (CHR ''(''), (CHR ''E''), (CHR ''L''), (CHR ''2''), (CHR '')''), (CHR '' ''), (CHR ''&''), (CHR ''&''), (CHR '' ''), (CHR ''!''), (CHR ''(''), (CHR ''I''), (CHR ''s''), (CHR ''S''), (CHR ''e''), (CHR ''c''), (CHR ''u''), (CHR ''r''), (CHR ''e''), (CHR ''(''), (CHR '')''), (CHR '')''), (CHR '')''), (CHR '' ''), (CHR ''&''), (CHR ''&''), (CHR '' ''), (CHR ''(''), (CHR ''(''), (CHR ''(''), (CHR ''H''), (CHR ''C''), (CHR ''R''), (CHR ''_''), (CHR ''E''), (CHR ''L''), (CHR ''2''), (CHR '')''), (CHR ''.''), (CHR ''T''), (CHR ''G''), (CHR ''E''), (CHR '' ''), (CHR ''=''), (CHR ''=''), (CHR '' ''), (char_of_nat 39), (CHR ''1''), (char_of_nat 39), (CHR '')''), (CHR '' ''), (CHR ''|''), (CHR ''|''), (CHR '' ''), (CHR ''(''), (CHR ''(''), (CHR ''M''), (CHR ''D''), (CHR ''C''), (CHR ''R''), (CHR ''_''), (CHR ''E''), (CHR ''L''), (CHR ''2''), (CHR '')''), (CHR ''.''), (CHR ''T''), (CHR ''D''), (CHR ''E''), (CHR '' ''), (CHR ''=''), (CHR ''=''), (CHR '' ''), (char_of_nat 39), (CHR ''1''), (char_of_nat 39), (CHR '')''), (CHR '')''), (CHR '')'')]) \<then> + (assert_exp w__6 ([(CHR ''(''), (CHR ''(''), (CHR ''H''), (CHR ''a''), (CHR ''v''), (CHR ''e''), (CHR ''E''), (CHR ''L''), (CHR ''(''), (CHR ''E''), (CHR ''L''), (CHR ''2''), (CHR '')''), (CHR '' ''), (CHR ''&''), (CHR ''&''), (CHR '' ''), (CHR ''!''), (CHR ''(''), (CHR ''I''), (CHR ''s''), (CHR ''S''), (CHR ''e''), (CHR ''c''), (CHR ''u''), (CHR ''r''), (CHR ''e''), (CHR ''(''), (CHR '')''), (CHR '')''), (CHR '')''), (CHR '' ''), (CHR ''&''), (CHR ''&''), (CHR '' ''), (CHR ''(''), (CHR ''(''), (CHR ''(''), (CHR ''H''), (CHR ''C''), (CHR ''R''), (CHR ''_''), (CHR ''E''), (CHR ''L''), (CHR ''2''), (CHR '')''), (CHR ''.''), (CHR ''T''), (CHR ''G''), (CHR ''E''), (CHR '' ''), (CHR ''=''), (CHR ''=''), (CHR '' ''), (CHR 0x27), (CHR ''1''), (CHR 0x27), (CHR '')''), (CHR '' ''), (CHR ''|''), (CHR ''|''), (CHR '' ''), (CHR ''(''), (CHR ''(''), (CHR ''M''), (CHR ''D''), (CHR ''C''), (CHR ''R''), (CHR ''_''), (CHR ''E''), (CHR ''L''), (CHR ''2''), (CHR '')''), (CHR ''.''), (CHR ''T''), (CHR ''D''), (CHR ''E''), (CHR '' ''), (CHR ''=''), (CHR ''=''), (CHR '' ''), (CHR 0x27), (CHR ''1''), (CHR 0x27), (CHR '')''), (CHR '')''), (CHR '')'')]) \<then> (ThisInstrAddr (( 64 :: int)::ii) () :: ( 64 Word.word) M)) \<bind> (\<lambda> (preferred_exception_return :: 64 bits) . (let (vect_offset :: ii) = ((( 0 :: int)::ii)) in (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (vaddress :: 64 bits) . @@ -17404,7 +17412,7 @@ definition AArch32_EnterMode :: "(5)Word.word \<Rightarrow>(32)Word.word \<Righ (GetPSRFromPSTATE () :: ( 32 Word.word) M)) \<bind> (\<lambda> (spsr :: 32 bits) . read_reg PSTATE_ref \<bind> (\<lambda> (w__3 :: ProcState) . (((((if ((((ProcState_M w__3) = M32_Monitor))) then - (read_reg SCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . + (read_reg SCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . write_reg SCR_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__4 (( 0 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 32 Word.word))) @@ -17644,42 +17652,42 @@ definition aarch64_float_convert_int :: " int \<Rightarrow>('fltsize::len)itsel (let intsize = (size_itself_int intsize) in (let fltsize = (size_itself_int fltsize) in (CheckFPAdvSIMDEnabled64 () \<then> - (undefined_bitvector fltsize :: (( 'fltsize::len)Word.word) M)) \<bind> (\<lambda> (fltval :: 'fltsize bits) . - (undefined_bitvector intsize :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (intval :: 'intsize bits) . + (undefined_bitvector fltsize :: (( 'fltsize::len)Word.word) M)) \<bind> (\<lambda> fltval . + (undefined_bitvector intsize :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> intval . (case op1 of FPConvOp_CVT_FtoI => - (aget_V fltsize n :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'fltsize bits) . + (aget_V fltsize n :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'fltsize::len)Word.word) . (let fltval = w__0 in - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . - (FPToFixed intsize fltval (( 0 :: int)::ii) unsigned w__1 rounding :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'intsize - bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . + (FPToFixed intsize fltval (( 0 :: int)::ii) unsigned w__1 rounding :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: + ( 'intsize::len)Word.word) . (let intval = w__2 in aset_X d intval))))) | FPConvOp_CVT_ItoF => - (aget_X intsize n :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: 'intsize bits) . + (aget_X intsize n :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'intsize::len)Word.word) . (let intval = w__3 in - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . - (FixedToFP fltsize intval (( 0 :: int)::ii) unsigned w__4 rounding :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: 'fltsize - bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . + (FixedToFP fltsize intval (( 0 :: int)::ii) unsigned w__4 rounding :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + ( 'fltsize::len)Word.word) . (let fltval = w__5 in aset_V d fltval))))) | FPConvOp_MOV_FtoI => - (aget_Vpart fltsize n part :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: 'fltsize bits) . + (aget_Vpart fltsize n part :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: ( 'fltsize::len)Word.word) . (let fltval = w__6 in - (ZeroExtend__0 fltval ((make_the_value intsize :: ( 'intsize::len)itself)) :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: 'intsize - bits) . + (ZeroExtend__0 fltval ((make_the_value intsize )) :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: + ( 'intsize::len)Word.word) . (let intval = w__7 in aset_X d intval)))) | FPConvOp_MOV_ItoF => - (aget_X intsize n :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__8 :: 'intsize bits) . + (aget_X intsize n :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__8 :: ( 'intsize::len)Word.word) . (let intval = w__8 in (let fltval = ((slice intval (( 0 :: int)::ii) fltsize :: ( 'fltsize::len)Word.word)) in aset_Vpart d part fltval))) | FPConvOp_CVT_FtoI_JS => - (aget_V fltsize n :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__9 :: 'fltsize bits) . + (aget_V fltsize n :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__9 :: ( 'fltsize::len)Word.word) . (let fltval = w__9 in - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . - (FPToFixedJS intsize fltval w__10 True :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__11 :: 'intsize bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . + (FPToFixedJS intsize fltval w__10 True :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__11 :: ( 'intsize::len)Word.word) . (let intval = w__11 in (ZeroExtend__0 ((slice intval (( 0 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) ((make_the_value (( 64 :: int)::ii) :: 64 itself)) @@ -17695,10 +17703,10 @@ definition aarch64_float_convert_fp :: " int \<Rightarrow>('dstsize::len)itself (let srcsize = (size_itself_int srcsize) in (let dstsize = (size_itself_int dstsize) in (CheckFPAdvSIMDEnabled64 () \<then> - (undefined_bitvector dstsize :: (( 'dstsize::len)Word.word) M)) \<bind> (\<lambda> (result :: 'dstsize bits) . - (aget_V srcsize n :: (( 'srcsize::len)Word.word) M) \<bind> (\<lambda> (operand :: 'srcsize bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . - (FPConvert__1 dstsize operand w__0 :: (( 'dstsize::len)Word.word) M) \<bind> (\<lambda> (w__1 :: 'dstsize bits) . + (undefined_bitvector dstsize :: (( 'dstsize::len)Word.word) M)) \<bind> (\<lambda> result . + (aget_V srcsize n :: (( 'srcsize::len)Word.word) M) \<bind> (\<lambda> operand . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . + (FPConvert__1 dstsize operand w__0 :: (( 'dstsize::len)Word.word) M) \<bind> (\<lambda> (w__1 :: ( 'dstsize::len)Word.word) . (let result = w__1 in aset_V d result))))))))" @@ -17710,23 +17718,23 @@ definition aarch64_float_convert_fix :: " int \<Rightarrow>('fltsize::len)itsel (let intsize = (size_itself_int intsize) in (let fltsize = (size_itself_int fltsize) in (CheckFPAdvSIMDEnabled64 () \<then> - (undefined_bitvector fltsize :: (( 'fltsize::len)Word.word) M)) \<bind> (\<lambda> (fltval :: 'fltsize bits) . - (undefined_bitvector intsize :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (intval :: 'intsize bits) . + (undefined_bitvector fltsize :: (( 'fltsize::len)Word.word) M)) \<bind> (\<lambda> fltval . + (undefined_bitvector intsize :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> intval . (case op1 of FPConvOp_CVT_FtoI => - (aget_V fltsize n :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__0 :: 'fltsize bits) . + (aget_V fltsize n :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'fltsize::len)Word.word) . (let fltval = w__0 in - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . - (FPToFixed intsize fltval fracbits unsigned w__1 rounding :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'intsize - bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . + (FPToFixed intsize fltval fracbits unsigned w__1 rounding :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: + ( 'intsize::len)Word.word) . (let intval = w__2 in aset_X d intval))))) | FPConvOp_CVT_ItoF => - (aget_X intsize n :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: 'intsize bits) . + (aget_X intsize n :: (( 'intsize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'intsize::len)Word.word) . (let intval = w__3 in - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . - (FixedToFP fltsize intval fracbits unsigned w__4 rounding :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: 'fltsize - bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . + (FixedToFP fltsize intval fracbits unsigned w__4 rounding :: (( 'fltsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + ( 'fltsize::len)Word.word) . (let fltval = w__5 in aset_V d fltval))))) ))))))" @@ -17744,7 +17752,7 @@ definition aarch64_float_compare_uncond :: " bool \<Rightarrow> int \<Rightarro (aget_V (( 8 :: int)::ii) n :: ( 8 Word.word) M)) \<bind> (\<lambda> (operand1 :: 8 bits) . (if cmp_with_zero then (FPZero (( 8 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: ( 8 Word.word) M) else (aget_V (( 8 :: int)::ii) m :: ( 8 Word.word) M)) \<bind> (\<lambda> (operand2 :: 8 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__2 :: ( 4 Word.word) M) \<bind> (\<lambda> split_vec . (let (tup__0, tup__1, tup__2, tup__3) = ((subrange_vec_dec split_vec (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word), @@ -17767,7 +17775,7 @@ definition aarch64_float_compare_uncond :: " bool \<Rightarrow> int \<Rightarro (aget_V (( 16 :: int)::ii) n :: ( 16 Word.word) M)) \<bind> (\<lambda> (operand1 :: 16 bits) . (if cmp_with_zero then (FPZero (( 16 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: ( 16 Word.word) M) else (aget_V (( 16 :: int)::ii) m :: ( 16 Word.word) M)) \<bind> (\<lambda> (operand2 :: 16 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__9 :: ( 4 Word.word) M) \<bind> (\<lambda> split_vec . (let (tup__0, tup__1, tup__2, tup__3) = ((subrange_vec_dec split_vec (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word), @@ -17790,7 +17798,7 @@ definition aarch64_float_compare_uncond :: " bool \<Rightarrow> int \<Rightarro (aget_V (( 32 :: int)::ii) n :: ( 32 Word.word) M)) \<bind> (\<lambda> (operand1 :: 32 bits) . (if cmp_with_zero then (FPZero (( 32 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: ( 32 Word.word) M) else (aget_V (( 32 :: int)::ii) m :: ( 32 Word.word) M)) \<bind> (\<lambda> (operand2 :: 32 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__16 :: ( 4 Word.word) M) \<bind> (\<lambda> split_vec . (let (tup__0, tup__1, tup__2, tup__3) = ((subrange_vec_dec split_vec (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word), @@ -17813,7 +17821,7 @@ definition aarch64_float_compare_uncond :: " bool \<Rightarrow> int \<Rightarro (aget_V (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (operand1 :: 64 bits) . (if cmp_with_zero then (FPZero (( 64 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: ( 64 Word.word) M) else (aget_V (( 64 :: int)::ii) m :: ( 64 Word.word) M)) \<bind> (\<lambda> (operand2 :: 64 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__23 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__23 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__23 :: ( 4 Word.word) M) \<bind> (\<lambda> split_vec . (let (tup__0, tup__1, tup__2, tup__3) = ((subrange_vec_dec split_vec (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word), @@ -17836,7 +17844,7 @@ definition aarch64_float_compare_uncond :: " bool \<Rightarrow> int \<Rightarro (aget_V (( 128 :: int)::ii) n :: ( 128 Word.word) M)) \<bind> (\<lambda> (operand1 :: 128 bits) . (if cmp_with_zero then (FPZero (( 128 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: ( 128 Word.word) M) else (aget_V (( 128 :: int)::ii) m :: ( 128 Word.word) M)) \<bind> (\<lambda> (operand2 :: 128 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__30 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__30 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__30 :: ( 4 Word.word) M) \<bind> (\<lambda> split_vec . (let (tup__0, tup__1, tup__2, tup__3) = ((subrange_vec_dec split_vec (( 3 :: int)::ii) (( 3 :: int)::ii) :: 1 Word.word), @@ -17870,7 +17878,7 @@ definition aarch64_float_compare_cond :: "(4)Word.word \<Rightarrow> int \<Righ (aget_V (( 8 :: int)::ii) m :: ( 8 Word.word) M) \<bind> (\<lambda> (operand2 :: 8 bits) . ConditionHolds condition \<bind> (\<lambda> (w__0 :: bool) . (if w__0 then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__1 :: ( 4 Word.word) M)) else return flags) \<bind> (\<lambda> (flags :: 4 Word.word) . (let (tup__0, tup__1, tup__2, tup__3) = @@ -17896,7 +17904,7 @@ definition aarch64_float_compare_cond :: "(4)Word.word \<Rightarrow> int \<Righ (aget_V (( 16 :: int)::ii) m :: ( 16 Word.word) M) \<bind> (\<lambda> (operand2 :: 16 bits) . ConditionHolds condition \<bind> (\<lambda> (w__7 :: bool) . (if w__7 then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__8 :: ( 4 Word.word) M)) else return flags) \<bind> (\<lambda> (flags :: 4 Word.word) . (let (tup__0, tup__1, tup__2, tup__3) = @@ -17922,7 +17930,7 @@ definition aarch64_float_compare_cond :: "(4)Word.word \<Rightarrow> int \<Righ (aget_V (( 32 :: int)::ii) m :: ( 32 Word.word) M) \<bind> (\<lambda> (operand2 :: 32 bits) . ConditionHolds condition \<bind> (\<lambda> (w__14 :: bool) . (if w__14 then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__15 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__15 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__15 :: ( 4 Word.word) M)) else return flags) \<bind> (\<lambda> (flags :: 4 Word.word) . (let (tup__0, tup__1, tup__2, tup__3) = @@ -17948,7 +17956,7 @@ definition aarch64_float_compare_cond :: "(4)Word.word \<Rightarrow> int \<Righ (aget_V (( 64 :: int)::ii) m :: ( 64 Word.word) M) \<bind> (\<lambda> (operand2 :: 64 bits) . ConditionHolds condition \<bind> (\<lambda> (w__21 :: bool) . (if w__21 then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__22 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__22 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__22 :: ( 4 Word.word) M)) else return flags) \<bind> (\<lambda> (flags :: 4 Word.word) . (let (tup__0, tup__1, tup__2, tup__3) = @@ -17974,7 +17982,7 @@ definition aarch64_float_compare_cond :: "(4)Word.word \<Rightarrow> int \<Righ (aget_V (( 128 :: int)::ii) m :: ( 128 Word.word) M) \<bind> (\<lambda> (operand2 :: 128 bits) . ConditionHolds condition \<bind> (\<lambda> (w__28 :: bool) . (if w__28 then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__29 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__29 :: 32 bits) . (FPCompare operand1 operand2 signal_all_nans w__29 :: ( 4 Word.word) M)) else return flags) \<bind> (\<lambda> (flags :: 4 Word.word) . (let (tup__0, tup__1, tup__2, tup__3) = @@ -18011,7 +18019,7 @@ definition aarch64_float_arithmetic_unary :: " int \<Rightarrow> int \<Rightarr | FPUnaryOp_ABS => (FPAbs operand :: ( 8 Word.word) M) | FPUnaryOp_NEG => (FPNeg operand :: ( 8 Word.word) M) | FPUnaryOp_SQRT => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . (FPSqrt operand w__2 :: ( 8 Word.word) M)) ) \<bind> (\<lambda> (result :: 8 bits) . aset_V d result)))) @@ -18027,7 +18035,7 @@ definition aarch64_float_arithmetic_unary :: " int \<Rightarrow> int \<Rightarr | FPUnaryOp_ABS => (FPAbs operand :: ( 16 Word.word) M) | FPUnaryOp_NEG => (FPNeg operand :: ( 16 Word.word) M) | FPUnaryOp_SQRT => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . (FPSqrt operand w__6 :: ( 16 Word.word) M)) ) \<bind> (\<lambda> (result :: 16 bits) . aset_V d result)))) @@ -18043,7 +18051,7 @@ definition aarch64_float_arithmetic_unary :: " int \<Rightarrow> int \<Rightarr | FPUnaryOp_ABS => (FPAbs operand :: ( 32 Word.word) M) | FPUnaryOp_NEG => (FPNeg operand :: ( 32 Word.word) M) | FPUnaryOp_SQRT => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . (FPSqrt operand w__10 :: ( 32 Word.word) M)) ) \<bind> (\<lambda> (result :: 32 bits) . aset_V d result)))) @@ -18059,7 +18067,7 @@ definition aarch64_float_arithmetic_unary :: " int \<Rightarrow> int \<Rightarr | FPUnaryOp_ABS => (FPAbs operand :: ( 64 Word.word) M) | FPUnaryOp_NEG => (FPNeg operand :: ( 64 Word.word) M) | FPUnaryOp_SQRT => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 bits) . (FPSqrt operand w__14 :: ( 64 Word.word) M)) ) \<bind> (\<lambda> (result :: 64 bits) . aset_V d result)))) @@ -18075,7 +18083,7 @@ definition aarch64_float_arithmetic_unary :: " int \<Rightarrow> int \<Rightarr | FPUnaryOp_ABS => (FPAbs operand :: ( 128 Word.word) M) | FPUnaryOp_NEG => (FPNeg operand :: ( 128 Word.word) M) | FPUnaryOp_SQRT => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 bits) . (FPSqrt operand w__18 :: ( 128 Word.word) M)) ) \<bind> (\<lambda> (result :: 128 bits) . aset_V d result)))) @@ -18095,8 +18103,8 @@ definition aarch64_float_arithmetic_round :: " int \<Rightarrow> int \<Rightarr CheckFPAdvSIMDEnabled64 () ) \<then> (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M)) \<bind> (\<lambda> (result :: 8 bits) . (aget_V (( 8 :: int)::ii) n :: ( 8 Word.word) M) \<bind> (\<lambda> (operand :: 8 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . - (FPRoundInt operand w__0 rounding exact :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . + (FPRoundInt operand w__0 rounding exact :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 Word.word) . (let result = w__1 in aset_V d result)))))) else if (((l__113 = (( 16 :: int)::ii)))) then @@ -18106,8 +18114,8 @@ definition aarch64_float_arithmetic_round :: " int \<Rightarrow> int \<Rightarr CheckFPAdvSIMDEnabled64 () ) \<then> (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M)) \<bind> (\<lambda> (result :: 16 bits) . (aget_V (( 16 :: int)::ii) n :: ( 16 Word.word) M) \<bind> (\<lambda> (operand :: 16 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . - (FPRoundInt operand w__2 rounding exact :: ( 16 Word.word) M) \<bind> (\<lambda> (w__3 :: 16 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . + (FPRoundInt operand w__2 rounding exact :: ( 16 Word.word) M) \<bind> (\<lambda> (w__3 :: 16 Word.word) . (let result = w__3 in aset_V d result)))))) else if (((l__113 = (( 32 :: int)::ii)))) then @@ -18117,8 +18125,8 @@ definition aarch64_float_arithmetic_round :: " int \<Rightarrow> int \<Rightarr CheckFPAdvSIMDEnabled64 () ) \<then> (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (result :: 32 bits) . (aget_V (( 32 :: int)::ii) n :: ( 32 Word.word) M) \<bind> (\<lambda> (operand :: 32 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . - (FPRoundInt operand w__4 rounding exact :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . + (FPRoundInt operand w__4 rounding exact :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 Word.word) . (let result = w__5 in aset_V d result)))))) else if (((l__113 = (( 64 :: int)::ii)))) then @@ -18128,8 +18136,8 @@ definition aarch64_float_arithmetic_round :: " int \<Rightarrow> int \<Rightarr CheckFPAdvSIMDEnabled64 () ) \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (result :: 64 bits) . (aget_V (( 64 :: int)::ii) n :: ( 64 Word.word) M) \<bind> (\<lambda> (operand :: 64 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . - (FPRoundInt operand w__6 rounding exact :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . + (FPRoundInt operand w__6 rounding exact :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 Word.word) . (let result = w__7 in aset_V d result)))))) else if (((l__113 = (( 128 :: int)::ii)))) then @@ -18139,8 +18147,8 @@ definition aarch64_float_arithmetic_round :: " int \<Rightarrow> int \<Rightarr CheckFPAdvSIMDEnabled64 () ) \<then> (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M)) \<bind> (\<lambda> (result :: 128 bits) . (aget_V (( 128 :: int)::ii) n :: ( 128 Word.word) M) \<bind> (\<lambda> (operand :: 128 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 Word.word) . - (FPRoundInt operand w__8 rounding exact :: ( 128 Word.word) M) \<bind> (\<lambda> (w__9 :: 128 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 bits) . + (FPRoundInt operand w__8 rounding exact :: ( 128 Word.word) M) \<bind> (\<lambda> (w__9 :: 128 Word.word) . (let result = w__9 in aset_V d result)))))) else @@ -18160,8 +18168,8 @@ definition aarch64_float_arithmetic_mul_product :: " int \<Rightarrow> int \<Ri (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M)) \<bind> (\<lambda> (result :: 8 bits) . (aget_V (( 8 :: int)::ii) n :: ( 8 Word.word) M) \<bind> (\<lambda> (operand1 :: 8 bits) . (aget_V (( 8 :: int)::ii) m :: ( 8 Word.word) M) \<bind> (\<lambda> (operand2 :: 8 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . - (FPMul operand1 operand2 w__0 :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . + (FPMul operand1 operand2 w__0 :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 Word.word) . (let result = w__1 in (if negated then (FPNeg result :: ( 8 Word.word) M) else return result) \<bind> (\<lambda> (result :: 8 bits) . @@ -18174,8 +18182,8 @@ definition aarch64_float_arithmetic_mul_product :: " int \<Rightarrow> int \<Ri (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M)) \<bind> (\<lambda> (result :: 16 bits) . (aget_V (( 16 :: int)::ii) n :: ( 16 Word.word) M) \<bind> (\<lambda> (operand1 :: 16 bits) . (aget_V (( 16 :: int)::ii) m :: ( 16 Word.word) M) \<bind> (\<lambda> (operand2 :: 16 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__3 :: 32 Word.word) . - (FPMul operand1 operand2 w__3 :: ( 16 Word.word) M) \<bind> (\<lambda> (w__4 :: 16 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__3 :: 32 bits) . + (FPMul operand1 operand2 w__3 :: ( 16 Word.word) M) \<bind> (\<lambda> (w__4 :: 16 Word.word) . (let result = w__4 in (if negated then (FPNeg result :: ( 16 Word.word) M) else return result) \<bind> (\<lambda> (result :: 16 bits) . @@ -18188,8 +18196,8 @@ definition aarch64_float_arithmetic_mul_product :: " int \<Rightarrow> int \<Ri (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (result :: 32 bits) . (aget_V (( 32 :: int)::ii) n :: ( 32 Word.word) M) \<bind> (\<lambda> (operand1 :: 32 bits) . (aget_V (( 32 :: int)::ii) m :: ( 32 Word.word) M) \<bind> (\<lambda> (operand2 :: 32 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . - (FPMul operand1 operand2 w__6 :: ( 32 Word.word) M) \<bind> (\<lambda> (w__7 :: 32 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . + (FPMul operand1 operand2 w__6 :: ( 32 Word.word) M) \<bind> (\<lambda> (w__7 :: 32 Word.word) . (let result = w__7 in (if negated then (FPNeg result :: ( 32 Word.word) M) else return result) \<bind> (\<lambda> (result :: 32 bits) . @@ -18202,8 +18210,8 @@ definition aarch64_float_arithmetic_mul_product :: " int \<Rightarrow> int \<Ri (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (result :: 64 bits) . (aget_V (( 64 :: int)::ii) n :: ( 64 Word.word) M) \<bind> (\<lambda> (operand1 :: 64 bits) . (aget_V (( 64 :: int)::ii) m :: ( 64 Word.word) M) \<bind> (\<lambda> (operand2 :: 64 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 Word.word) . - (FPMul operand1 operand2 w__9 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__10 :: 64 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 bits) . + (FPMul operand1 operand2 w__9 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__10 :: 64 Word.word) . (let result = w__10 in (if negated then (FPNeg result :: ( 64 Word.word) M) else return result) \<bind> (\<lambda> (result :: 64 bits) . @@ -18216,8 +18224,8 @@ definition aarch64_float_arithmetic_mul_product :: " int \<Rightarrow> int \<Ri (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M)) \<bind> (\<lambda> (result :: 128 bits) . (aget_V (( 128 :: int)::ii) n :: ( 128 Word.word) M) \<bind> (\<lambda> (operand1 :: 128 bits) . (aget_V (( 128 :: int)::ii) m :: ( 128 Word.word) M) \<bind> (\<lambda> (operand2 :: 128 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 Word.word) . - (FPMul operand1 operand2 w__12 :: ( 128 Word.word) M) \<bind> (\<lambda> (w__13 :: 128 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 bits) . + (FPMul operand1 operand2 w__12 :: ( 128 Word.word) M) \<bind> (\<lambda> (w__13 :: 128 Word.word) . (let result = w__13 in (if negated then (FPNeg result :: ( 128 Word.word) M) else return result) \<bind> (\<lambda> (result :: 128 bits) . @@ -18244,8 +18252,8 @@ definition aarch64_float_arithmetic_mul_addsub :: " int \<Rightarrow> int \<Rig else return operanda) \<bind> (\<lambda> (operanda :: 8 bits) . (if op1_neg then (FPNeg operand1 :: ( 8 Word.word) M) else return operand1) \<bind> (\<lambda> (operand1 :: 8 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . - (FPMulAdd operanda operand1 operand2 w__2 :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: 8 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . + (FPMulAdd operanda operand1 operand2 w__2 :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: 8 Word.word) . (let result = w__3 in aset_V d result)))))))))) else if (((l__103 = (( 16 :: int)::ii)))) then @@ -18261,8 +18269,8 @@ definition aarch64_float_arithmetic_mul_addsub :: " int \<Rightarrow> int \<Rig else return operanda) \<bind> (\<lambda> (operanda :: 16 bits) . (if op1_neg then (FPNeg operand1 :: ( 16 Word.word) M) else return operand1) \<bind> (\<lambda> (operand1 :: 16 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . - (FPMulAdd operanda operand1 operand2 w__6 :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . + (FPMulAdd operanda operand1 operand2 w__6 :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 Word.word) . (let result = w__7 in aset_V d result)))))))))) else if (((l__103 = (( 32 :: int)::ii)))) then @@ -18278,8 +18286,8 @@ definition aarch64_float_arithmetic_mul_addsub :: " int \<Rightarrow> int \<Rig else return operanda) \<bind> (\<lambda> (operanda :: 32 bits) . (if op1_neg then (FPNeg operand1 :: ( 32 Word.word) M) else return operand1) \<bind> (\<lambda> (operand1 :: 32 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . - (FPMulAdd operanda operand1 operand2 w__10 :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: 32 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . + (FPMulAdd operanda operand1 operand2 w__10 :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: 32 Word.word) . (let result = w__11 in aset_V d result)))))))))) else if (((l__103 = (( 64 :: int)::ii)))) then @@ -18295,8 +18303,8 @@ definition aarch64_float_arithmetic_mul_addsub :: " int \<Rightarrow> int \<Rig else return operanda) \<bind> (\<lambda> (operanda :: 64 bits) . (if op1_neg then (FPNeg operand1 :: ( 64 Word.word) M) else return operand1) \<bind> (\<lambda> (operand1 :: 64 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 Word.word) . - (FPMulAdd operanda operand1 operand2 w__14 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 bits) . + (FPMulAdd operanda operand1 operand2 w__14 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 Word.word) . (let result = w__15 in aset_V d result)))))))))) else if (((l__103 = (( 128 :: int)::ii)))) then @@ -18312,8 +18320,8 @@ definition aarch64_float_arithmetic_mul_addsub :: " int \<Rightarrow> int \<Rig else return operanda) \<bind> (\<lambda> (operanda :: 128 bits) . (if op1_neg then (FPNeg operand1 :: ( 128 Word.word) M) else return operand1) \<bind> (\<lambda> (operand1 :: 128 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 Word.word) . - (FPMulAdd operanda operand1 operand2 w__18 :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: 128 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 bits) . + (FPMulAdd operanda operand1 operand2 w__18 :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: 128 Word.word) . (let result = w__19 in aset_V d result)))))))))) else @@ -18335,16 +18343,16 @@ definition aarch64_float_arithmetic_maxmin :: " int \<Rightarrow> int \<Rightar (aget_V (( 8 :: int)::ii) m :: ( 8 Word.word) M) \<bind> (\<lambda> (operand2 :: 8 bits) . (case operation of FPMaxMinOp_MAX => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . (FPMax operand1 operand2 w__0 :: ( 8 Word.word) M)) | FPMaxMinOp_MIN => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . (FPMin operand1 operand2 w__2 :: ( 8 Word.word) M)) | FPMaxMinOp_MAXNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . (FPMaxNum operand1 operand2 w__4 :: ( 8 Word.word) M)) | FPMaxMinOp_MINNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . (FPMinNum operand1 operand2 w__6 :: ( 8 Word.word) M)) ) \<bind> (\<lambda> (result :: 8 bits) . aset_V d result))))) @@ -18358,16 +18366,16 @@ definition aarch64_float_arithmetic_maxmin :: " int \<Rightarrow> int \<Rightar (aget_V (( 16 :: int)::ii) m :: ( 16 Word.word) M) \<bind> (\<lambda> (operand2 :: 16 bits) . (case operation of FPMaxMinOp_MAX => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 bits) . (FPMax operand1 operand2 w__8 :: ( 16 Word.word) M)) | FPMaxMinOp_MIN => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . (FPMin operand1 operand2 w__10 :: ( 16 Word.word) M)) | FPMaxMinOp_MAXNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 bits) . (FPMaxNum operand1 operand2 w__12 :: ( 16 Word.word) M)) | FPMaxMinOp_MINNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 bits) . (FPMinNum operand1 operand2 w__14 :: ( 16 Word.word) M)) ) \<bind> (\<lambda> (result :: 16 bits) . aset_V d result))))) @@ -18381,16 +18389,16 @@ definition aarch64_float_arithmetic_maxmin :: " int \<Rightarrow> int \<Rightar (aget_V (( 32 :: int)::ii) m :: ( 32 Word.word) M) \<bind> (\<lambda> (operand2 :: 32 bits) . (case operation of FPMaxMinOp_MAX => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: 32 bits) . (FPMax operand1 operand2 w__16 :: ( 32 Word.word) M)) | FPMaxMinOp_MIN => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 bits) . (FPMin operand1 operand2 w__18 :: ( 32 Word.word) M)) | FPMaxMinOp_MAXNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 bits) . (FPMaxNum operand1 operand2 w__20 :: ( 32 Word.word) M)) | FPMaxMinOp_MINNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__22 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__22 :: 32 bits) . (FPMinNum operand1 operand2 w__22 :: ( 32 Word.word) M)) ) \<bind> (\<lambda> (result :: 32 bits) . aset_V d result))))) @@ -18404,16 +18412,16 @@ definition aarch64_float_arithmetic_maxmin :: " int \<Rightarrow> int \<Rightar (aget_V (( 64 :: int)::ii) m :: ( 64 Word.word) M) \<bind> (\<lambda> (operand2 :: 64 bits) . (case operation of FPMaxMinOp_MAX => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__24 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__24 :: 32 bits) . (FPMax operand1 operand2 w__24 :: ( 64 Word.word) M)) | FPMaxMinOp_MIN => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__26 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__26 :: 32 bits) . (FPMin operand1 operand2 w__26 :: ( 64 Word.word) M)) | FPMaxMinOp_MAXNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__28 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__28 :: 32 bits) . (FPMaxNum operand1 operand2 w__28 :: ( 64 Word.word) M)) | FPMaxMinOp_MINNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__30 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__30 :: 32 bits) . (FPMinNum operand1 operand2 w__30 :: ( 64 Word.word) M)) ) \<bind> (\<lambda> (result :: 64 bits) . aset_V d result))))) @@ -18427,16 +18435,16 @@ definition aarch64_float_arithmetic_maxmin :: " int \<Rightarrow> int \<Rightar (aget_V (( 128 :: int)::ii) m :: ( 128 Word.word) M) \<bind> (\<lambda> (operand2 :: 128 bits) . (case operation of FPMaxMinOp_MAX => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__32 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__32 :: 32 bits) . (FPMax operand1 operand2 w__32 :: ( 128 Word.word) M)) | FPMaxMinOp_MIN => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__34 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__34 :: 32 bits) . (FPMin operand1 operand2 w__34 :: ( 128 Word.word) M)) | FPMaxMinOp_MAXNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__36 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__36 :: 32 bits) . (FPMaxNum operand1 operand2 w__36 :: ( 128 Word.word) M)) | FPMaxMinOp_MINNUM => - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__38 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__38 :: 32 bits) . (FPMinNum operand1 operand2 w__38 :: ( 128 Word.word) M)) ) \<bind> (\<lambda> (result :: 128 bits) . aset_V d result))))) @@ -18457,8 +18465,8 @@ definition aarch64_float_arithmetic_div :: " int \<Rightarrow> int \<Rightarrow (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M)) \<bind> (\<lambda> (result :: 8 bits) . (aget_V (( 8 :: int)::ii) n :: ( 8 Word.word) M) \<bind> (\<lambda> (operand1 :: 8 bits) . (aget_V (( 8 :: int)::ii) m :: ( 8 Word.word) M) \<bind> (\<lambda> (operand2 :: 8 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . - (FPDiv operand1 operand2 w__0 :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . + (FPDiv operand1 operand2 w__0 :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 Word.word) . (let result = w__1 in aset_V d result))))))) else if (((l__93 = (( 16 :: int)::ii)))) then @@ -18469,8 +18477,8 @@ definition aarch64_float_arithmetic_div :: " int \<Rightarrow> int \<Rightarrow (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M)) \<bind> (\<lambda> (result :: 16 bits) . (aget_V (( 16 :: int)::ii) n :: ( 16 Word.word) M) \<bind> (\<lambda> (operand1 :: 16 bits) . (aget_V (( 16 :: int)::ii) m :: ( 16 Word.word) M) \<bind> (\<lambda> (operand2 :: 16 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . - (FPDiv operand1 operand2 w__2 :: ( 16 Word.word) M) \<bind> (\<lambda> (w__3 :: 16 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . + (FPDiv operand1 operand2 w__2 :: ( 16 Word.word) M) \<bind> (\<lambda> (w__3 :: 16 Word.word) . (let result = w__3 in aset_V d result))))))) else if (((l__93 = (( 32 :: int)::ii)))) then @@ -18481,8 +18489,8 @@ definition aarch64_float_arithmetic_div :: " int \<Rightarrow> int \<Rightarrow (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (result :: 32 bits) . (aget_V (( 32 :: int)::ii) n :: ( 32 Word.word) M) \<bind> (\<lambda> (operand1 :: 32 bits) . (aget_V (( 32 :: int)::ii) m :: ( 32 Word.word) M) \<bind> (\<lambda> (operand2 :: 32 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . - (FPDiv operand1 operand2 w__4 :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . + (FPDiv operand1 operand2 w__4 :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 Word.word) . (let result = w__5 in aset_V d result))))))) else if (((l__93 = (( 64 :: int)::ii)))) then @@ -18493,8 +18501,8 @@ definition aarch64_float_arithmetic_div :: " int \<Rightarrow> int \<Rightarrow (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (result :: 64 bits) . (aget_V (( 64 :: int)::ii) n :: ( 64 Word.word) M) \<bind> (\<lambda> (operand1 :: 64 bits) . (aget_V (( 64 :: int)::ii) m :: ( 64 Word.word) M) \<bind> (\<lambda> (operand2 :: 64 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . - (FPDiv operand1 operand2 w__6 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . + (FPDiv operand1 operand2 w__6 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 Word.word) . (let result = w__7 in aset_V d result))))))) else if (((l__93 = (( 128 :: int)::ii)))) then @@ -18505,8 +18513,8 @@ definition aarch64_float_arithmetic_div :: " int \<Rightarrow> int \<Rightarrow (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M)) \<bind> (\<lambda> (result :: 128 bits) . (aget_V (( 128 :: int)::ii) n :: ( 128 Word.word) M) \<bind> (\<lambda> (operand1 :: 128 bits) . (aget_V (( 128 :: int)::ii) m :: ( 128 Word.word) M) \<bind> (\<lambda> (operand2 :: 128 bits) . - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 Word.word) . - (FPDiv operand1 operand2 w__8 :: ( 128 Word.word) M) \<bind> (\<lambda> (w__9 :: 128 bits) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 bits) . + (FPDiv operand1 operand2 w__8 :: ( 128 Word.word) M) \<bind> (\<lambda> (w__9 :: 128 Word.word) . (let result = w__9 in aset_V d result))))))) else @@ -18527,10 +18535,10 @@ definition aarch64_float_arithmetic_addsub :: " int \<Rightarrow> int \<Rightar (aget_V (( 8 :: int)::ii) n :: ( 8 Word.word) M) \<bind> (\<lambda> (operand1 :: 8 bits) . (aget_V (( 8 :: int)::ii) m :: ( 8 Word.word) M) \<bind> (\<lambda> (operand2 :: 8 bits) . (if sub_op then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . (FPSub operand1 operand2 w__0 :: ( 8 Word.word) M)) else - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . (FPAdd operand1 operand2 w__2 :: ( 8 Word.word) M))) \<bind> (\<lambda> (result :: 8 bits) . aset_V d result))))) else if (((l__88 = (( 16 :: int)::ii)))) then @@ -18542,10 +18550,10 @@ definition aarch64_float_arithmetic_addsub :: " int \<Rightarrow> int \<Rightar (aget_V (( 16 :: int)::ii) n :: ( 16 Word.word) M) \<bind> (\<lambda> (operand1 :: 16 bits) . (aget_V (( 16 :: int)::ii) m :: ( 16 Word.word) M) \<bind> (\<lambda> (operand2 :: 16 bits) . (if sub_op then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . (FPSub operand1 operand2 w__4 :: ( 16 Word.word) M)) else - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . (FPAdd operand1 operand2 w__6 :: ( 16 Word.word) M))) \<bind> (\<lambda> (result :: 16 bits) . aset_V d result))))) else if (((l__88 = (( 32 :: int)::ii)))) then @@ -18557,10 +18565,10 @@ definition aarch64_float_arithmetic_addsub :: " int \<Rightarrow> int \<Rightar (aget_V (( 32 :: int)::ii) n :: ( 32 Word.word) M) \<bind> (\<lambda> (operand1 :: 32 bits) . (aget_V (( 32 :: int)::ii) m :: ( 32 Word.word) M) \<bind> (\<lambda> (operand2 :: 32 bits) . (if sub_op then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 bits) . (FPSub operand1 operand2 w__8 :: ( 32 Word.word) M)) else - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . (FPAdd operand1 operand2 w__10 :: ( 32 Word.word) M))) \<bind> (\<lambda> (result :: 32 bits) . aset_V d result))))) else if (((l__88 = (( 64 :: int)::ii)))) then @@ -18572,10 +18580,10 @@ definition aarch64_float_arithmetic_addsub :: " int \<Rightarrow> int \<Rightar (aget_V (( 64 :: int)::ii) n :: ( 64 Word.word) M) \<bind> (\<lambda> (operand1 :: 64 bits) . (aget_V (( 64 :: int)::ii) m :: ( 64 Word.word) M) \<bind> (\<lambda> (operand2 :: 64 bits) . (if sub_op then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 bits) . (FPSub operand1 operand2 w__12 :: ( 64 Word.word) M)) else - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 bits) . (FPAdd operand1 operand2 w__14 :: ( 64 Word.word) M))) \<bind> (\<lambda> (result :: 64 bits) . aset_V d result))))) else if (((l__88 = (( 128 :: int)::ii)))) then @@ -18587,10 +18595,10 @@ definition aarch64_float_arithmetic_addsub :: " int \<Rightarrow> int \<Rightar (aget_V (( 128 :: int)::ii) n :: ( 128 Word.word) M) \<bind> (\<lambda> (operand1 :: 128 bits) . (aget_V (( 128 :: int)::ii) m :: ( 128 Word.word) M) \<bind> (\<lambda> (operand2 :: 128 bits) . (if sub_op then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: 32 bits) . (FPSub operand1 operand2 w__16 :: ( 128 Word.word) M)) else - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__18 :: 32 bits) . (FPAdd operand1 operand2 w__18 :: ( 128 Word.word) M))) \<bind> (\<lambda> (result :: 128 bits) . aset_V d result))))) else @@ -18793,7 +18801,7 @@ definition AArch64_CheckPermission :: " Permissions \<Rightarrow>(64)Word.word if fail1 then (let secondstage = False in (let s2fs1walk = False in - (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__27 :: 52 bits) . + (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M) \<bind> (\<lambda> (w__27 :: 52 Word.word) . (let ipaddress = w__27 in AArch64_PermissionFault ipaddress level acctype ((\<not> failedread)) secondstage s2fs1walk)))) else AArch64_NoFault () )))))))))))))))))))))))))))" @@ -18921,7 +18929,7 @@ definition AArch64_TranslateAddress :: "(64)Word.word \<Rightarrow> AccType \<R (let (result :: AddressDescriptor) = ((result (| AddressDescriptor_fault := w__0 |))) in return result)) else return result) \<bind> (\<lambda> (result :: AddressDescriptor) . - (ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . + (ZeroExtend__1 (( 64 :: int)::ii) vaddress :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (let (result :: AddressDescriptor) = ((result (| AddressDescriptor_vaddress := w__1 |))) in return result)))))" @@ -18980,9 +18988,9 @@ definition aset_Mem :: "(64)Word.word \<Rightarrow> int \<Rightarrow> AccType \ ((slice value_name (((( 8 :: int)::ii) * i)) (( 8 :: int)::ii) :: 8 Word.word))))) else if ((((((size1 = (( 16 :: int)::ii)))) \<and> ((((((acctype = AccType_VEC))) \<or> (((acctype = AccType_VECSTREAM))))))))) then AArch64_aset_MemSingle address (( 8 :: int)::ii) acctype aligned - ((slice value_name (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) \<then> + ((slice ((Word.ucast value_name :: 128 Word.word)) (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) \<then> AArch64_aset_MemSingle ((add_vec_int address (( 8 :: int)::ii) :: 64 Word.word)) (( 8 :: int)::ii) acctype aligned - ((slice value_name (( 64 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) + ((slice ((Word.ucast value_name :: 128 Word.word)) (( 64 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) else AArch64_aset_MemSingle address size1 acctype aligned value_name))))))))))" @@ -19044,12 +19052,18 @@ definition aget_Mem :: "(64)Word.word \<Rightarrow> int \<Rightarrow> AccType \ (AArch64_aget_MemSingle address (( 8 :: int)::ii) acctype aligned :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . (let value_name = - ((set_slice (((( 8 :: int)::ii) * (( 16 :: int)::ii))) (( 64 :: int)::ii) value_name (( 0 :: int)::ii) w__2 + ((Word.ucast + ((set_slice (((( 8 :: int)::ii) * (( 16 :: int)::ii))) (( 64 :: int)::ii) + ((Word.ucast value_name :: 128 Word.word)) (( 0 :: int)::ii) w__2 + :: 128 Word.word)) :: ( 'p8_times_size_::len)Word.word)) in (AArch64_aget_MemSingle ((add_vec_int address (( 8 :: int)::ii) :: 64 Word.word)) (( 8 :: int)::ii) acctype aligned :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . (let value_name = - ((set_slice (((( 8 :: int)::ii) * (( 16 :: int)::ii))) (( 64 :: int)::ii) value_name (( 64 :: int)::ii) w__3 + ((Word.ucast + ((set_slice (((( 8 :: int)::ii) * (( 16 :: int)::ii))) (( 64 :: int)::ii) + ((Word.ucast value_name :: 128 Word.word)) (( 64 :: int)::ii) w__3 + :: 128 Word.word)) :: ( 'p8_times_size_::len)Word.word)) in return value_name)))) else (AArch64_aget_MemSingle address size1 acctype aligned :: (( 'p8_times_size_::len)Word.word) M)) \<bind> (\<lambda> value_name . @@ -19070,7 +19084,7 @@ definition aarch64_memory_vector_single_nowb :: "('datasize::len)itself \<Right (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (offs :: 64 bits) . (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (rval :: 128 bits) . - (undefined_bitvector esize :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> (element :: 'esize bits) . + (undefined_bitvector esize :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> element . undefined_int () \<bind> (\<lambda> (s :: ii) . (let (ebytes :: int) = (ex_int ((esize div (( 8 :: int)::ii)))) in (assert_exp True ('''') \<then> @@ -19080,31 +19094,31 @@ definition aarch64_memory_vector_single_nowb :: "('datasize::len)itself \<Right (if replicate1 then (foreachM (index_list (( 0 :: int)::ii) ((selem - (( 1 :: int)::ii))) (( 1 :: int)::ii)) (element, offs, t) (\<lambda> s varstup . (let (element, offs, t) = varstup in - (aget_Mem ((add_vec address offs :: 64 Word.word)) ebytes AccType_VEC :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'esize - bits) . + (aget_Mem ((add_vec address offs :: 64 Word.word)) ebytes AccType_VEC :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: + ( 'esize::len)Word.word) . (let element = w__2 in (let (v :: int) = (ex_int ((datasize div esize))) in (assert_exp True ('''') \<then> aset_V t ((replicate_bits element v :: ( 'datasize::len)Word.word))) \<then> ((let (offs :: 64 bits) = ((add_vec_int offs ebytes :: 64 Word.word)) in (let (t :: ii) = (((((ex_int t)) + (( 1 :: int)::ii))) mod (( 32 :: int)::ii)) in - return (element, offs, t)))))))))) \<bind> (\<lambda> varstup . (let ((element :: 'esize bits), (offs :: 64 - bits), (t :: ii)) = varstup in + return (element, offs, t)))))))))) \<bind> (\<lambda> varstup . (let (element, (offs :: 64 bits), (t :: ii)) = varstup in return offs)) else (foreachM (index_list (( 0 :: int)::ii) ((selem - (( 1 :: int)::ii))) (( 1 :: int)::ii)) (offs, rval, t) (\<lambda> s varstup . (let (offs, rval, t) = varstup in - (aget_V (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__3 :: 128 bits) . + (aget_V (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__3 :: 128 Word.word) . (let rval = w__3 in (if (((memop = MemOp_LOAD))) then (aget_Mem ((add_vec address offs :: 64 Word.word)) ebytes AccType_VEC :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: ( 'esize::len)Word.word) . - (aset_Elem__0 rval index1 ((make_the_value esize :: ( 'esize::len)itself)) w__4 - :: ( 128 Word.word) M) \<bind> (\<lambda> (w__5 :: 128 bits) . + (aset_Elem__0 rval index1 ((make_the_value esize )) w__4 :: ( 128 Word.word) M) \<bind> (\<lambda> (w__5 :: + 128 Word.word) . (let rval = w__5 in aset_V t rval \<then> return rval))) else - (aget_Elem__0 rval index1 ((make_the_value esize :: ( 'esize::len)itself)) :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> w__6 . + (aget_Elem__0 rval index1 ((make_the_value esize )) :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + ( 'esize::len)Word.word) . aset_Mem ((add_vec address offs :: 64 Word.word)) ebytes AccType_VEC w__6 \<then> return rval)) \<bind> (\<lambda> (rval :: 128 bits) . (let (offs :: 64 bits) = ((add_vec_int offs ebytes :: 64 Word.word)) in @@ -19130,7 +19144,7 @@ definition aarch64_memory_vector_multiple_nowb :: "('datasize::len)itself \<Rig CheckFPAdvSIMDEnabled64 () ) \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (offs :: 64 bits) . - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (rval :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> rval . undefined_int () \<bind> (\<lambda> (e :: ii) . undefined_int () \<bind> (\<lambda> (r :: ii) . undefined_int () \<bind> (\<lambda> (s :: ii) . @@ -19147,23 +19161,24 @@ definition aarch64_memory_vector_multiple_nowb :: "('datasize::len)itself \<Rig (let tt = (((t + r)) mod (( 32 :: int)::ii)) in (foreachM (index_list (( 0 :: int)::ii) ((selem - (( 1 :: int)::ii))) (( 1 :: int)::ii)) (offs, rval, tt) (\<lambda> s varstup . (let (offs, rval, tt) = varstup in - (aget_V datasize tt :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'datasize bits) . + (aget_V datasize tt :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: ( 'datasize::len)Word.word) . (let rval = w__2 in (if (((memop = MemOp_LOAD))) then (aget_Mem ((add_vec address offs :: 64 Word.word)) ebytes AccType_VEC :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'esize::len)Word.word) . - (aset_Elem__0 rval e ((make_the_value esize :: ( 'esize::len)itself)) w__3 - :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: 'datasize bits) . + (aset_Elem__0 rval e ((make_the_value esize )) w__3 :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: + ( 'datasize::len)Word.word) . (let rval = w__4 in aset_V tt rval \<then> return rval))) else - (aget_Elem__0 rval e ((make_the_value esize :: ( 'esize::len)itself)) :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> w__5 . + (aget_Elem__0 rval e ((make_the_value esize )) :: (( 'esize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + ( 'esize::len)Word.word) . aset_Mem ((add_vec address offs :: 64 Word.word)) ebytes AccType_VEC w__5 \<then> - return rval)) \<bind> (\<lambda> (rval :: 'datasize bits) . + return rval)) \<bind> (\<lambda> rval . (let (offs :: 64 bits) = ((add_vec_int offs ebytes :: 64 Word.word)) in (let (tt :: ii) = (((((ex_int tt)) + (( 1 :: int)::ii))) mod (( 32 :: int)::ii)) in - return (offs, rval, tt)))))))))))))))) \<bind> (\<lambda> varstup . (let ((offs :: 64 bits), (rval :: 'datasize - bits), (tt :: ii)) = varstup in + return (offs, rval, tt)))))))))))))))) \<bind> (\<lambda> varstup . (let ((offs :: 64 bits), rval, (tt :: + ii)) = varstup in if wback then (if (((m \<noteq> (( 31 :: int)::ii)))) then (aget_X (( 64 :: int)::ii) m :: ( 64 Word.word) M) else return offs) \<bind> (\<lambda> (offs :: 64 bits) . @@ -19191,12 +19206,12 @@ definition aarch64_memory_single_simdfp_register :: " AccType \<Rightarrow> int else address) in (case memop of MemOp_STORE => - (aget_V (( 8 :: int)::ii) t :: ( 8 Word.word) M) \<bind> (\<lambda> (w__2 :: 8 bits) . + (aget_V (( 8 :: int)::ii) t :: ( 8 Word.word) M) \<bind> (\<lambda> (w__2 :: 8 Word.word) . (let data = w__2 in aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: + 8 Word.word) . (let data = w__3 in aset_V t data)) ) \<then> @@ -19222,12 +19237,12 @@ definition aarch64_memory_single_simdfp_register :: " AccType \<Rightarrow> int else address) in (case memop of MemOp_STORE => - (aget_V (( 16 :: int)::ii) t :: ( 16 Word.word) M) \<bind> (\<lambda> (w__6 :: 16 bits) . + (aget_V (( 16 :: int)::ii) t :: ( 16 Word.word) M) \<bind> (\<lambda> (w__6 :: 16 Word.word) . (let data = w__6 in aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: + 16 Word.word) . (let data = w__7 in aset_V t data)) ) \<then> @@ -19253,12 +19268,12 @@ definition aarch64_memory_single_simdfp_register :: " AccType \<Rightarrow> int else address) in (case memop of MemOp_STORE => - (aget_V (( 32 :: int)::ii) t :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . + (aget_V (( 32 :: int)::ii) t :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . (let data = w__10 in aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: + 32 Word.word) . (let data = w__11 in aset_V t data)) ) \<then> @@ -19284,12 +19299,12 @@ definition aarch64_memory_single_simdfp_register :: " AccType \<Rightarrow> int else address) in (case memop of MemOp_STORE => - (aget_V (( 64 :: int)::ii) t :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: 64 bits) . + (aget_V (( 64 :: int)::ii) t :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: 64 Word.word) . (let data = w__14 in aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: + 64 Word.word) . (let data = w__15 in aset_V t data)) ) \<then> @@ -19315,12 +19330,12 @@ definition aarch64_memory_single_simdfp_register :: " AccType \<Rightarrow> int else address) in (case memop of MemOp_STORE => - (aget_V (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__18 :: 128 bits) . + (aget_V (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__18 :: 128 Word.word) . (let data = w__18 in aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: + 128 Word.word) . (let data = w__19 in aset_V t data)) ) \<then> @@ -19354,12 +19369,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_postidx :: " AccType \ else address) in (case memop of MemOp_STORE => - (aget_V (( 8 :: int)::ii) t :: ( 8 Word.word) M) \<bind> (\<lambda> (w__2 :: 8 bits) . + (aget_V (( 8 :: int)::ii) t :: ( 8 Word.word) M) \<bind> (\<lambda> (w__2 :: 8 Word.word) . (let data = w__2 in aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: + 8 Word.word) . (let data = w__3 in aset_V t data)) ) \<then> @@ -19384,12 +19399,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_postidx :: " AccType \ else address) in (case memop of MemOp_STORE => - (aget_V (( 16 :: int)::ii) t :: ( 16 Word.word) M) \<bind> (\<lambda> (w__6 :: 16 bits) . + (aget_V (( 16 :: int)::ii) t :: ( 16 Word.word) M) \<bind> (\<lambda> (w__6 :: 16 Word.word) . (let data = w__6 in aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: + 16 Word.word) . (let data = w__7 in aset_V t data)) ) \<then> @@ -19414,12 +19429,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_postidx :: " AccType \ else address) in (case memop of MemOp_STORE => - (aget_V (( 32 :: int)::ii) t :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . + (aget_V (( 32 :: int)::ii) t :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . (let data = w__10 in aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: + 32 Word.word) . (let data = w__11 in aset_V t data)) ) \<then> @@ -19444,12 +19459,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_postidx :: " AccType \ else address) in (case memop of MemOp_STORE => - (aget_V (( 64 :: int)::ii) t :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: 64 bits) . + (aget_V (( 64 :: int)::ii) t :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: 64 Word.word) . (let data = w__14 in aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: + 64 Word.word) . (let data = w__15 in aset_V t data)) ) \<then> @@ -19474,12 +19489,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_postidx :: " AccType \ else address) in (case memop of MemOp_STORE => - (aget_V (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__18 :: 128 bits) . + (aget_V (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__18 :: 128 Word.word) . (let data = w__18 in aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: + 128 Word.word) . (let data = w__19 in aset_V t data)) ) \<then> @@ -19513,12 +19528,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_offset_normal :: " Acc else address) in (case memop of MemOp_STORE => - (aget_V (( 8 :: int)::ii) t :: ( 8 Word.word) M) \<bind> (\<lambda> (w__2 :: 8 bits) . + (aget_V (( 8 :: int)::ii) t :: ( 8 Word.word) M) \<bind> (\<lambda> (w__2 :: 8 Word.word) . (let data = w__2 in aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: + 8 Word.word) . (let data = w__3 in aset_V t data)) ) \<then> @@ -19543,12 +19558,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_offset_normal :: " Acc else address) in (case memop of MemOp_STORE => - (aget_V (( 16 :: int)::ii) t :: ( 16 Word.word) M) \<bind> (\<lambda> (w__6 :: 16 bits) . + (aget_V (( 16 :: int)::ii) t :: ( 16 Word.word) M) \<bind> (\<lambda> (w__6 :: 16 Word.word) . (let data = w__6 in aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: + 16 Word.word) . (let data = w__7 in aset_V t data)) ) \<then> @@ -19573,12 +19588,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_offset_normal :: " Acc else address) in (case memop of MemOp_STORE => - (aget_V (( 32 :: int)::ii) t :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . + (aget_V (( 32 :: int)::ii) t :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . (let data = w__10 in aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: + 32 Word.word) . (let data = w__11 in aset_V t data)) ) \<then> @@ -19603,12 +19618,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_offset_normal :: " Acc else address) in (case memop of MemOp_STORE => - (aget_V (( 64 :: int)::ii) t :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: 64 bits) . + (aget_V (( 64 :: int)::ii) t :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: 64 Word.word) . (let data = w__14 in aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: + 64 Word.word) . (let data = w__15 in aset_V t data)) ) \<then> @@ -19633,12 +19648,12 @@ definition aarch64_memory_single_simdfp_immediate_signed_offset_normal :: " Acc else address) in (case memop of MemOp_STORE => - (aget_V (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__18 :: 128 bits) . + (aget_V (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__18 :: 128 Word.word) . (let data = w__18 in aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype data)) | MemOp_LOAD => - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: + 128 Word.word) . (let data = w__19 in aset_V t data)) ) \<then> @@ -19662,20 +19677,20 @@ definition aarch64_memory_ordered :: " AccType \<Rightarrow>('datasize::len)its (let datasize = (size_itself_int datasize) in (assert_exp True (''datasize constraint'') \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data . (let dbytes = (ex_int ((datasize div (( 8 :: int)::ii)))) in (assert_exp True ('''') \<then> (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M))) \<bind> (\<lambda> (address :: 64 bits) . (case memop of MemOp_STORE => - (aget_X ((int (size data))) t :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'datasize bits) . + (aget_X ((int (size data))) t :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: ( 'datasize::len)Word.word) . (let data = w__2 in aset_Mem address dbytes acctype data)) | MemOp_LOAD => - (aget_Mem address dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: 'datasize bits) . + (aget_Mem address dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'datasize::len)Word.word) . (let data = w__3 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: ( 'regsize::len)Word.word) . aset_X t w__4))) ))))))))" @@ -19763,14 +19778,14 @@ definition aarch64_memory_orderedrcpc :: " AccType \<Rightarrow>('datasize::len (let datasize = (size_itself_int datasize) in (assert_exp True (''datasize constraint'') \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data . (let dbytes = (ex_int ((datasize div (( 8 :: int)::ii)))) in (assert_exp True ('''') \<then> (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M))) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'datasize bits) . + (aget_Mem address dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: ( 'datasize::len)Word.word) . (let data = w__2 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'regsize::len)Word.word) . aset_X t w__3))))))))))" @@ -19836,7 +19851,7 @@ definition aarch64_memory_literal_simdfp :: "(64)Word.word \<Rightarrow> int \< (let (address :: 64 bits) = ((add_vec w__0 offset :: 64 Word.word)) in (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (data :: 32 bits) . (CheckFPAdvSIMDEnabled64 () \<then> - (aget_Mem address (( 4 :: int)::ii) AccType_VEC :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__1 :: 32 bits) . + (aget_Mem address (( 4 :: int)::ii) AccType_VEC :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__1 :: 32 Word.word) . (let data = w__1 in aset_V t data))))) else if (((l__70 = (( 8 :: int)::ii)))) then @@ -19846,7 +19861,7 @@ definition aarch64_memory_literal_simdfp :: "(64)Word.word \<Rightarrow> int \< (let (address :: 64 bits) = ((add_vec w__2 offset :: 64 Word.word)) in (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (data :: 64 bits) . (CheckFPAdvSIMDEnabled64 () \<then> - (aget_Mem address (( 8 :: int)::ii) AccType_VEC :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__3 :: 64 bits) . + (aget_Mem address (( 8 :: int)::ii) AccType_VEC :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__3 :: 64 Word.word) . (let data = w__3 in aset_V t data))))) else if (((l__70 = (( 16 :: int)::ii)))) then @@ -19856,7 +19871,7 @@ definition aarch64_memory_literal_simdfp :: "(64)Word.word \<Rightarrow> int \< (let (address :: 64 bits) = ((add_vec w__4 offset :: 64 Word.word)) in (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (data :: 128 bits) . (CheckFPAdvSIMDEnabled64 () \<then> - (aget_Mem address (( 16 :: int)::ii) AccType_VEC :: ( 128 Word.word) M)) \<bind> (\<lambda> (w__5 :: 128 bits) . + (aget_Mem address (( 16 :: int)::ii) AccType_VEC :: ( 128 Word.word) M)) \<bind> (\<lambda> (w__5 :: 128 Word.word) . (let data = w__5 in aset_V t data))))) else assert_exp True ('''') \<then> assert_exp True (''''))" @@ -19869,13 +19884,13 @@ definition aarch64_memory_literal_general :: " MemOp \<Rightarrow>(64)Word.word (let size1 = (size_itself_int size1) in (aget_PC () :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . (let (address :: 64 bits) = ((add_vec w__0 offset :: 64 Word.word)) in - (undefined_bitvector size1 :: (( 'size::len)Word.word) M) \<bind> (\<lambda> (data :: 'size bits) . + (undefined_bitvector size1 :: (( 'size::len)Word.word) M) \<bind> (\<lambda> data . (case memop of MemOp_LOAD => assert_exp True ('''') \<then> ((let bytes = (size1 div (( 8 :: int)::ii)) in (assert_exp True ('''') \<then> - (aget_Mem address bytes AccType_NORMAL :: (( 'size::len)Word.word) M)) \<bind> (\<lambda> (w__1 :: 'size bits) . + (aget_Mem address bytes AccType_NORMAL :: (( 'size::len)Word.word) M)) \<bind> (\<lambda> (w__1 :: ( 'size::len)Word.word) . (let data = w__1 in if signed then (SignExtend__0 data ((make_the_value (( 64 :: int)::ii) :: 64 itself)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: @@ -19902,7 +19917,7 @@ definition memory_literal_general_decode :: "(2)Word.word \<Rightarrow>(1)Word. (let size1 = ((( 4 :: int)::ii)) in (SignExtend__0 ((concat_vec imm19 (vec_of_bits [B0,B0] :: 2 Word.word) :: 21 Word.word)) ((make_the_value (( 64 :: int)::ii) :: 64 itself)) - :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . + :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . (let offset = w__0 in aarch64_memory_literal_general MemOp_LOAD offset False ((make_the_value (( 32 :: int)::ii) :: 32 itself)) t))))))))) @@ -19916,7 +19931,7 @@ definition memory_literal_general_decode :: "(2)Word.word \<Rightarrow>(1)Word. (let size1 = ((( 8 :: int)::ii)) in (SignExtend__0 ((concat_vec imm19 (vec_of_bits [B0,B0] :: 2 Word.word) :: 21 Word.word)) ((make_the_value (( 64 :: int)::ii) :: 64 itself)) - :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . + :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (let offset = w__1 in aarch64_memory_literal_general MemOp_LOAD offset False ((make_the_value (( 64 :: int)::ii) :: 64 itself)) t))))))))) @@ -19931,7 +19946,7 @@ definition memory_literal_general_decode :: "(2)Word.word \<Rightarrow>(1)Word. (let signed = True in (SignExtend__0 ((concat_vec imm19 (vec_of_bits [B0,B0] :: 2 Word.word) :: 21 Word.word)) ((make_the_value (( 64 :: int)::ii) :: 64 itself)) - :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . + :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . (let offset = w__2 in aarch64_memory_literal_general MemOp_LOAD offset True ((make_the_value (( 32 :: int)::ii) :: 32 itself)) t)))))))))) @@ -19945,7 +19960,7 @@ definition memory_literal_general_decode :: "(2)Word.word \<Rightarrow>(1)Word. (let memop = MemOp_PREFETCH in (SignExtend__0 ((concat_vec imm19 (vec_of_bits [B0,B0] :: 2 Word.word) :: 21 Word.word)) ((make_the_value (( 64 :: int)::ii) :: 64 itself)) - :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 bits) . + :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . (let offset = w__3 in aarch64_memory_literal_general MemOp_PREFETCH offset False ((make_the_value (((( 8 :: int)::ii) * (( 32 :: int)::ii))) :: 256 itself)) t))))))))))" @@ -19965,12 +19980,12 @@ definition aarch64_memory_atomicops_swp :: " int \<Rightarrow> AccType \<Righta (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (data :: 8 bits) . (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__2 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__2 :: + 8 Word.word) . (let data = w__2 in (aget_X ((int (size data))) s :: ( 8 Word.word) M) \<bind> (\<lambda> w__3 . (aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) stacctype w__3 \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__4 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__4 :: ( 'regsize::len)Word.word) . aset_X t w__4)))))))))) else if (((l__65 = (( 16 :: int)::ii)))) then @@ -19983,12 +19998,12 @@ definition aarch64_memory_atomicops_swp :: " int \<Rightarrow> AccType \<Righta (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (data :: 16 bits) . (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: + 16 Word.word) . (let data = w__7 in (aget_X ((int (size data))) s :: ( 16 Word.word) M) \<bind> (\<lambda> w__8 . (aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) stacctype w__8 \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__9 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__9 :: ( 'regsize::len)Word.word) . aset_X t w__9)))))))))) else if (((l__65 = (( 32 :: int)::ii)))) then @@ -20001,12 +20016,12 @@ definition aarch64_memory_atomicops_swp :: " int \<Rightarrow> AccType \<Righta (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (data :: 32 bits) . (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: + 32 Word.word) . (let data = w__12 in (aget_X ((int (size data))) s :: ( 32 Word.word) M) \<bind> (\<lambda> w__13 . (aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) stacctype w__13 \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14)))))))))) else if (((l__65 = (( 64 :: int)::ii)))) then @@ -20019,12 +20034,12 @@ definition aarch64_memory_atomicops_swp :: " int \<Rightarrow> AccType \<Righta (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (data :: 64 bits) . (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: + 64 Word.word) . (let data = w__17 in (aget_X ((int (size address))) s :: ( 64 Word.word) M) \<bind> (\<lambda> w__18 . (aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) stacctype w__18 \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__19 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__19 :: ( 'regsize::len)Word.word) . aset_X t w__19)))))))))) else if (((l__65 = (( 128 :: int)::ii)))) then @@ -20037,12 +20052,12 @@ definition aarch64_memory_atomicops_swp :: " int \<Rightarrow> AccType \<Righta (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (data :: 128 bits) . (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__22 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__22 :: + 128 Word.word) . (let data = w__22 in (aget_X ((int (size data))) s :: ( 128 Word.word) M) \<bind> (\<lambda> w__23 . (aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) stacctype w__23 \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__24 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__24 :: ( 'regsize::len)Word.word) . aset_X t w__24)))))))))) else @@ -20064,12 +20079,12 @@ definition aarch64_memory_atomicops_st :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (value_name :: 8 bits) . (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (data :: 8 bits) . (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (result :: 8 bits) . - (aget_X (( 8 :: int)::ii) s :: ( 8 Word.word) M) \<bind> (\<lambda> (w__0 :: 8 bits) . + (aget_X (( 8 :: int)::ii) s :: ( 8 Word.word) M) \<bind> (\<lambda> (w__0 :: 8 Word.word) . (let value_name = w__0 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: + 8 Word.word) . (let data = w__3 in (let (result :: 8 bits) = ((case op1 of @@ -20091,12 +20106,12 @@ definition aarch64_memory_atomicops_st :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (value_name :: 16 bits) . (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (data :: 16 bits) . (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (result :: 16 bits) . - (aget_X (( 16 :: int)::ii) s :: ( 16 Word.word) M) \<bind> (\<lambda> (w__4 :: 16 bits) . + (aget_X (( 16 :: int)::ii) s :: ( 16 Word.word) M) \<bind> (\<lambda> (w__4 :: 16 Word.word) . (let value_name = w__4 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: + 16 Word.word) . (let data = w__7 in (let (result :: 16 bits) = ((case op1 of @@ -20118,12 +20133,12 @@ definition aarch64_memory_atomicops_st :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (value_name :: 32 bits) . (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (data :: 32 bits) . (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (result :: 32 bits) . - (aget_X (( 32 :: int)::ii) s :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 bits) . + (aget_X (( 32 :: int)::ii) s :: ( 32 Word.word) M) \<bind> (\<lambda> (w__8 :: 32 Word.word) . (let value_name = w__8 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__11 :: + 32 Word.word) . (let data = w__11 in (let (result :: 32 bits) = ((case op1 of @@ -20145,12 +20160,12 @@ definition aarch64_memory_atomicops_st :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (value_name :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (data :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (result :: 64 bits) . - (aget_X (( 64 :: int)::ii) s :: ( 64 Word.word) M) \<bind> (\<lambda> (w__12 :: 64 bits) . + (aget_X (( 64 :: int)::ii) s :: ( 64 Word.word) M) \<bind> (\<lambda> (w__12 :: 64 Word.word) . (let value_name = w__12 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: + 64 Word.word) . (let data = w__15 in (let (result :: 64 bits) = ((case op1 of @@ -20172,12 +20187,12 @@ definition aarch64_memory_atomicops_st :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (value_name :: 128 bits) . (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (data :: 128 bits) . (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (result :: 128 bits) . - (aget_X (( 128 :: int)::ii) s :: ( 128 Word.word) M) \<bind> (\<lambda> (w__16 :: 128 bits) . + (aget_X (( 128 :: int)::ii) s :: ( 128 Word.word) M) \<bind> (\<lambda> (w__16 :: 128 Word.word) . (let value_name = w__16 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__19 :: + 128 Word.word) . (let data = w__19 in (let (result :: 128 bits) = ((case op1 of @@ -20210,12 +20225,12 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (value_name :: 8 bits) . (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (data :: 8 bits) . (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (result :: 8 bits) . - (aget_X (( 8 :: int)::ii) s :: ( 8 Word.word) M) \<bind> (\<lambda> (w__0 :: 8 bits) . + (aget_X (( 8 :: int)::ii) s :: ( 8 Word.word) M) \<bind> (\<lambda> (w__0 :: 8 Word.word) . (let value_name = w__0 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__3 :: + 8 Word.word) . (let data = w__3 in (let (result :: 8 bits) = ((case op1 of @@ -20229,7 +20244,7 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar | MemAtomicOp_UMIN => if ((((Word.uint data)) > ((Word.uint value_name)))) then value_name else data )) in (aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) stacctype result \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__4 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__4 :: ( 'regsize::len)Word.word) . aset_X t w__4)))))))))))))) else if (((l__55 = (( 16 :: int)::ii)))) then @@ -20242,12 +20257,12 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (value_name :: 16 bits) . (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (data :: 16 bits) . (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (result :: 16 bits) . - (aget_X (( 16 :: int)::ii) s :: ( 16 Word.word) M) \<bind> (\<lambda> (w__5 :: 16 bits) . + (aget_X (( 16 :: int)::ii) s :: ( 16 Word.word) M) \<bind> (\<lambda> (w__5 :: 16 Word.word) . (let value_name = w__5 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__8 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__8 :: + 16 Word.word) . (let data = w__8 in (let (result :: 16 bits) = ((case op1 of @@ -20261,7 +20276,7 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar | MemAtomicOp_UMIN => if ((((Word.uint data)) > ((Word.uint value_name)))) then value_name else data )) in (aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) stacctype result \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__9 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__9 :: ( 'regsize::len)Word.word) . aset_X t w__9)))))))))))))) else if (((l__55 = (( 32 :: int)::ii)))) then @@ -20274,12 +20289,12 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (value_name :: 32 bits) . (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (data :: 32 bits) . (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (result :: 32 bits) . - (aget_X (( 32 :: int)::ii) s :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 bits) . + (aget_X (( 32 :: int)::ii) s :: ( 32 Word.word) M) \<bind> (\<lambda> (w__10 :: 32 Word.word) . (let value_name = w__10 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__13 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__13 :: + 32 Word.word) . (let data = w__13 in (let (result :: 32 bits) = ((case op1 of @@ -20293,7 +20308,7 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar | MemAtomicOp_UMIN => if ((((Word.uint data)) > ((Word.uint value_name)))) then value_name else data )) in (aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) stacctype result \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14)))))))))))))) else if (((l__55 = (( 64 :: int)::ii)))) then @@ -20306,12 +20321,12 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (value_name :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (data :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (result :: 64 bits) . - (aget_X (( 64 :: int)::ii) s :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 bits) . + (aget_X (( 64 :: int)::ii) s :: ( 64 Word.word) M) \<bind> (\<lambda> (w__15 :: 64 Word.word) . (let value_name = w__15 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__18 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__18 :: + 64 Word.word) . (let data = w__18 in (let (result :: 64 bits) = ((case op1 of @@ -20325,7 +20340,7 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar | MemAtomicOp_UMIN => if ((((Word.uint data)) > ((Word.uint value_name)))) then value_name else data )) in (aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) stacctype result \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__19 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__19 :: ( 'regsize::len)Word.word) . aset_X t w__19)))))))))))))) else if (((l__55 = (( 128 :: int)::ii)))) then @@ -20338,12 +20353,12 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (value_name :: 128 bits) . (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (data :: 128 bits) . (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (result :: 128 bits) . - (aget_X (( 128 :: int)::ii) s :: ( 128 Word.word) M) \<bind> (\<lambda> (w__20 :: 128 bits) . + (aget_X (( 128 :: int)::ii) s :: ( 128 Word.word) M) \<bind> (\<lambda> (w__20 :: 128 Word.word) . (let value_name = w__20 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__23 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__23 :: + 128 Word.word) . (let data = w__23 in (let (result :: 128 bits) = ((case op1 of @@ -20357,7 +20372,7 @@ definition aarch64_memory_atomicops_ld :: " int \<Rightarrow> AccType \<Rightar | MemAtomicOp_UMIN => if ((((Word.uint data)) > ((Word.uint value_name)))) then value_name else data )) in (aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) stacctype result \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__24 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__24 :: ( 'regsize::len)Word.word) . aset_X t w__24)))))))))))))) else @@ -20381,19 +20396,19 @@ definition aarch64_memory_atomicops_cas_single :: " int \<Rightarrow> AccType \ (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (comparevalue :: 8 bits) . (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (newvalue :: 8 bits) . (undefined_bitvector (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (data :: 8 bits) . - (aget_X (( 8 :: int)::ii) s :: ( 8 Word.word) M) \<bind> (\<lambda> (w__0 :: 8 bits) . + (aget_X (( 8 :: int)::ii) s :: ( 8 Word.word) M) \<bind> (\<lambda> (w__0 :: 8 Word.word) . (let comparevalue = w__0 in - (aget_X (( 8 :: int)::ii) t :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 bits) . + (aget_X (( 8 :: int)::ii) t :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 Word.word) . (let newvalue = w__1 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: + 8 Word.word) . (let data = w__4 in ((if (((data = comparevalue))) then aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__5 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__5 :: ( 'regsize::len)Word.word) . aset_X s w__5))))))))))))))) else if (((l__50 = (( 16 :: int)::ii)))) then @@ -20406,19 +20421,19 @@ definition aarch64_memory_atomicops_cas_single :: " int \<Rightarrow> AccType \ (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (comparevalue :: 16 bits) . (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (newvalue :: 16 bits) . (undefined_bitvector (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (data :: 16 bits) . - (aget_X (( 16 :: int)::ii) s :: ( 16 Word.word) M) \<bind> (\<lambda> (w__6 :: 16 bits) . + (aget_X (( 16 :: int)::ii) s :: ( 16 Word.word) M) \<bind> (\<lambda> (w__6 :: 16 Word.word) . (let comparevalue = w__6 in - (aget_X (( 16 :: int)::ii) t :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 bits) . + (aget_X (( 16 :: int)::ii) t :: ( 16 Word.word) M) \<bind> (\<lambda> (w__7 :: 16 Word.word) . (let newvalue = w__7 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__10 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__10 :: + 16 Word.word) . (let data = w__10 in ((if (((data = comparevalue))) then aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__11 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__11 :: ( 'regsize::len)Word.word) . aset_X s w__11))))))))))))))) else if (((l__50 = (( 32 :: int)::ii)))) then @@ -20431,19 +20446,19 @@ definition aarch64_memory_atomicops_cas_single :: " int \<Rightarrow> AccType \ (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (comparevalue :: 32 bits) . (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (newvalue :: 32 bits) . (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (data :: 32 bits) . - (aget_X (( 32 :: int)::ii) s :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 bits) . + (aget_X (( 32 :: int)::ii) s :: ( 32 Word.word) M) \<bind> (\<lambda> (w__12 :: 32 Word.word) . (let comparevalue = w__12 in - (aget_X (( 32 :: int)::ii) t :: ( 32 Word.word) M) \<bind> (\<lambda> (w__13 :: 32 bits) . + (aget_X (( 32 :: int)::ii) t :: ( 32 Word.word) M) \<bind> (\<lambda> (w__13 :: 32 Word.word) . (let newvalue = w__13 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__16 :: + 32 Word.word) . (let data = w__16 in ((if (((data = comparevalue))) then aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__17 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__17 :: ( 'regsize::len)Word.word) . aset_X s w__17))))))))))))))) else if (((l__50 = (( 64 :: int)::ii)))) then @@ -20456,19 +20471,19 @@ definition aarch64_memory_atomicops_cas_single :: " int \<Rightarrow> AccType \ (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (comparevalue :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (newvalue :: 64 bits) . (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (data :: 64 bits) . - (aget_X (( 64 :: int)::ii) s :: ( 64 Word.word) M) \<bind> (\<lambda> (w__18 :: 64 bits) . + (aget_X (( 64 :: int)::ii) s :: ( 64 Word.word) M) \<bind> (\<lambda> (w__18 :: 64 Word.word) . (let comparevalue = w__18 in - (aget_X (( 64 :: int)::ii) t :: ( 64 Word.word) M) \<bind> (\<lambda> (w__19 :: 64 bits) . + (aget_X (( 64 :: int)::ii) t :: ( 64 Word.word) M) \<bind> (\<lambda> (w__19 :: 64 Word.word) . (let newvalue = w__19 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__22 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__22 :: + 64 Word.word) . (let data = w__22 in ((if (((data = comparevalue))) then aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__23 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__23 :: ( 'regsize::len)Word.word) . aset_X s w__23))))))))))))))) else if (((l__50 = (( 128 :: int)::ii)))) then @@ -20481,19 +20496,19 @@ definition aarch64_memory_atomicops_cas_single :: " int \<Rightarrow> AccType \ (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (comparevalue :: 128 bits) . (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (newvalue :: 128 bits) . (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M) \<bind> (\<lambda> (data :: 128 bits) . - (aget_X (( 128 :: int)::ii) s :: ( 128 Word.word) M) \<bind> (\<lambda> (w__24 :: 128 bits) . + (aget_X (( 128 :: int)::ii) s :: ( 128 Word.word) M) \<bind> (\<lambda> (w__24 :: 128 Word.word) . (let comparevalue = w__24 in - (aget_X (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__25 :: 128 bits) . + (aget_X (( 128 :: int)::ii) t :: ( 128 Word.word) M) \<bind> (\<lambda> (w__25 :: 128 Word.word) . (let newvalue = w__25 in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__28 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__28 :: + 128 Word.word) . (let data = w__28 in ((if (((data = comparevalue))) then aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__29 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X s w__29))))))))))))))) else @@ -20531,29 +20546,25 @@ definition aarch64_memory_atomicops_cas_pair :: " int \<Rightarrow> AccType \<R else (concat_vec t2 t1 :: 16 Word.word)) in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__4 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__4 :: + 16 Word.word) . (let data = w__4 in ((if (((data = comparevalue))) then aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> BigEndian () ) \<bind> (\<lambda> (w__5 :: bool) . if w__5 then - (ZeroExtend__0 ((slice data (( 8 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 8 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: ( 'regsize::len)Word.word) . (aset_X s w__6 \<then> - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__7 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__7)) else - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__8 :: ( 'regsize::len)Word.word) . (aset_X s w__8 \<then> - (ZeroExtend__0 ((slice data (( 8 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 8 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__9 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__9))))))))))))))))))))) else if (((l__45 = (( 16 :: int)::ii)))) then @@ -20580,29 +20591,25 @@ definition aarch64_memory_atomicops_cas_pair :: " int \<Rightarrow> AccType \<R else (concat_vec t2 t1 :: 32 Word.word)) in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__14 :: + 32 Word.word) . (let data = w__14 in ((if (((data = comparevalue))) then aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> BigEndian () ) \<bind> (\<lambda> (w__15 :: bool) . if w__15 then - (ZeroExtend__0 ((slice data (( 16 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 16 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__16 :: ( 'regsize::len)Word.word) . (aset_X s w__16 \<then> - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__17 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__17)) else - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__18 :: ( 'regsize::len)Word.word) . (aset_X s w__18 \<then> - (ZeroExtend__0 ((slice data (( 16 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 16 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__19 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__19))))))))))))))))))))) else if (((l__45 = (( 32 :: int)::ii)))) then @@ -20629,29 +20636,25 @@ definition aarch64_memory_atomicops_cas_pair :: " int \<Rightarrow> AccType \<R else (concat_vec t2 t1 :: 64 Word.word)) in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__24 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__24 :: + 64 Word.word) . (let data = w__24 in ((if (((data = comparevalue))) then aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> BigEndian () ) \<bind> (\<lambda> (w__25 :: bool) . if w__25 then - (ZeroExtend__0 ((slice data (( 32 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 32 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__26 :: ( 'regsize::len)Word.word) . (aset_X s w__26 \<then> - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__27 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__27)) else - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__28 :: ( 'regsize::len)Word.word) . (aset_X s w__28 \<then> - (ZeroExtend__0 ((slice data (( 32 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 32 :: int)::ii) (( 32 :: int)::ii) :: 32 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__29))))))))))))))))))))) else if (((l__45 = (( 64 :: int)::ii)))) then @@ -20678,29 +20681,25 @@ definition aarch64_memory_atomicops_cas_pair :: " int \<Rightarrow> AccType \<R else (concat_vec t2 t1 :: 128 Word.word)) in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__34 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__34 :: + 128 Word.word) . (let data = w__34 in ((if (((data = comparevalue))) then aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> BigEndian () ) \<bind> (\<lambda> (w__35 :: bool) . if w__35 then - (ZeroExtend__0 ((slice data (( 64 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 64 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__36 :: ( 'regsize::len)Word.word) . (aset_X s w__36 \<then> - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__37 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__37)) else - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: ( 'regsize::len)Word.word) . (aset_X s w__38 \<then> - (ZeroExtend__0 ((slice data (( 64 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 64 :: int)::ii) (( 64 :: int)::ii) :: 64 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__39 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__39))))))))))))))))))))) else if (((l__45 = (( 128 :: int)::ii)))) then @@ -20727,29 +20726,25 @@ definition aarch64_memory_atomicops_cas_pair :: " int \<Rightarrow> AccType \<R else (concat_vec t2 t1 :: 256 Word.word)) in (if (((n = (( 31 :: int)::ii)))) then CheckSPAlignment () \<then> (aget_SP (( 64 :: int)::ii) () :: ( 64 Word.word) M) else (aget_X (( 64 :: int)::ii) n :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (aget_Mem address (((( 256 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 256 Word.word) M) \<bind> (\<lambda> (w__44 :: 256 - bits) . + (aget_Mem address (((( 256 :: int)::ii) div (( 8 :: int)::ii))) ldacctype :: ( 256 Word.word) M) \<bind> (\<lambda> (w__44 :: + 256 Word.word) . (let data = w__44 in ((if (((data = comparevalue))) then aset_Mem address (((( 256 :: int)::ii) div (( 8 :: int)::ii))) stacctype newvalue else return () ) \<then> BigEndian () ) \<bind> (\<lambda> (w__45 :: bool) . if w__45 then - (ZeroExtend__0 ((slice data (( 128 :: int)::ii) (( 128 :: int)::ii) :: 128 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 128 :: int)::ii) (( 128 :: int)::ii) :: 128 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__46 :: ( 'regsize::len)Word.word) . (aset_X s w__46 \<then> - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 128 :: int)::ii) :: 128 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 128 :: int)::ii) :: 128 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__47 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__47)) else - (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 128 :: int)::ii) :: 128 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 0 :: int)::ii) (( 128 :: int)::ii) :: 128 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__48 :: ( 'regsize::len)Word.word) . (aset_X s w__48 \<then> - (ZeroExtend__0 ((slice data (( 128 :: int)::ii) (( 128 :: int)::ii) :: 128 Word.word)) - ((make_the_value regsize :: ( 'regsize::len)itself)) + (ZeroExtend__0 ((slice data (( 128 :: int)::ii) (( 128 :: int)::ii) :: 128 Word.word)) ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M)) \<bind> (\<lambda> (w__49 :: ( 'regsize::len)Word.word) . aset_X ((s + (( 1 :: int)::ii))) w__49))))))))))))))))))))) else @@ -20828,7 +20823,7 @@ definition AArch32_GenerateDebugExceptionsFrom :: "(2)Word.word \<Rightarrow> b (ELStateUsingAArch32 EL1 secure \<bind> (\<lambda> (w__0 :: bool) . return ((\<not> w__0)))) \<bind> (\<lambda> (w__1 :: bool) . if w__1 then - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__2 :: 1 bits) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M) \<bind> (\<lambda> (w__2 :: 1 Word.word) . (let mask1 = w__2 in AArch64_GenerateDebugExceptionsFrom from1 secure mask1)) else @@ -20914,7 +20909,7 @@ definition DebugExceptionReturnSS :: "(32)Word.word \<Rightarrow>((register_val else (let mask1 = ((vec_of_bits [access_vec_dec spsr (( 9 :: int)::ii)] :: 1 Word.word)) in AArch64_GenerateDebugExceptionsFrom dest secure mask1)) \<bind> (\<lambda> (enabled_at_dest :: bool) . - (DebugTargetFrom secure :: ( 2 Word.word) M) \<bind> (\<lambda> (w__17 :: 2 bits) . + (DebugTargetFrom secure :: ( 2 Word.word) M) \<bind> (\<lambda> (w__17 :: 2 Word.word) . (let ELd = w__17 in and_boolM (and_boolM (ELUsingAArch32 ELd \<bind> (\<lambda> (w__18 :: bool) . return ((\<not> w__18)))) @@ -20932,7 +20927,7 @@ definition SetPSTATEFromPSR :: "(32)Word.word \<Rightarrow>((register_value),(u " SetPSTATEFromPSR spsr__arg = ( (let spsr = spsr__arg in read_reg PSTATE_ref \<bind> (\<lambda> (w__0 :: ProcState) . - (DebugExceptionReturnSS spsr :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 bits) . + (DebugExceptionReturnSS spsr :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 Word.word) . (write_reg PSTATE_ref (w__0 (| ProcState_SS := w__1 |)) \<then> IllegalExceptionReturn spsr) \<bind> (\<lambda> (w__2 :: bool) . ((if w__2 then @@ -20988,7 +20983,7 @@ definition SetPSTATEFromPSR :: "(32)Word.word \<Rightarrow>((register_value),(u PSTATE_ref (w__17 (| ProcState_Q := ((vec_of_bits [access_vec_dec spsr (( 27 :: int)::ii)] :: 1 Word.word))|)) \<then> read_reg PSTATE_ref) \<bind> (\<lambda> (w__18 :: ProcState) . - (RestoredITBits spsr :: ( 8 Word.word) M) \<bind> (\<lambda> (w__19 :: 8 bits) . + (RestoredITBits spsr :: ( 8 Word.word) M) \<bind> (\<lambda> (w__19 :: 8 Word.word) . (write_reg PSTATE_ref (w__18 (| ProcState_IT := w__19 |)) \<then> read_reg PSTATE_ref) \<bind> (\<lambda> (w__20 :: ProcState) . (write_reg PSTATE_ref (w__20 (| ProcState_GE := ((slice spsr (( 16 :: int)::ii) (( 4 :: int)::ii) :: 4 Word.word))|)) \<then> @@ -21057,7 +21052,7 @@ definition DRPSInstruction :: " unit \<Rightarrow>((register_value),(unit),(exc (SetPSTATEFromPSR w__4 \<then> UsingAArch32 () ) \<bind> (\<lambda> (w__5 :: bool) . (if w__5 then - (undefined_bitvector (( 13 :: int)::ii) :: ( 13 Word.word) M) \<bind> (\<lambda> (w__6 :: 13 bits) . + (undefined_bitvector (( 13 :: int)::ii) :: ( 13 Word.word) M) \<bind> (\<lambda> (w__6 :: 13 Word.word) . (let split_vec = w__6 in (let (tup__0, tup__1, tup__2, tup__3, tup__4, tup__5, tup__6, tup__7, tup__8, tup__9) = ((subrange_vec_dec split_vec (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word), @@ -21096,12 +21091,12 @@ definition DRPSInstruction :: " unit \<Rightarrow>((register_value),(unit),(exc (w__17 (| ProcState_IT := ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))|)) \<then> read_reg PSTATE_ref) \<bind> (\<lambda> (w__18 :: ProcState) . (write_reg PSTATE_ref (w__18 (| ProcState_T := ((vec_of_bits [B1] :: 1 Word.word))|)) \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__19 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__19 :: 32 Word.word) . (write_reg DLR_ref w__19 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__20 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__20 :: 32 Word.word) . write_reg DSPSR_ref w__20))))))))))))))))) else - (undefined_bitvector (( 9 :: int)::ii) :: ( 9 Word.word) M) \<bind> (\<lambda> (w__21 :: 9 bits) . + (undefined_bitvector (( 9 :: int)::ii) :: ( 9 Word.word) M) \<bind> (\<lambda> (w__21 :: 9 Word.word) . (let split_vec = w__21 in (let (tup__0, tup__1, tup__2, tup__3, tup__4, tup__5, tup__6, tup__7, tup__8) = ((subrange_vec_dec split_vec (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word), @@ -21131,9 +21126,9 @@ definition DRPSInstruction :: " unit \<Rightarrow>((register_value),(unit),(exc (write_reg PSTATE_ref (w__29 (| ProcState_I := tup__7 |)) \<then> read_reg PSTATE_ref) \<bind> (\<lambda> (w__30 :: ProcState) . (write_reg PSTATE_ref (w__30 (| ProcState_F := tup__8 |)) \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__31 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__31 :: 64 Word.word) . (write_reg DLR_EL0_ref w__31 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__32 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__32 :: 32 Word.word) . write_reg DSPSR_EL0_ref w__32))))))))))))))) \<then> UpdateEDSCRFields () ))))))" @@ -21321,7 +21316,7 @@ definition UnallocatedEncoding :: " unit \<Rightarrow>((register_value),(unit), " UnallocatedEncoding _ = ( and_boolM ((UsingAArch32 () )) ((AArch32_ExecutingCP10or11Instr () )) \<bind> (\<lambda> (w__2 :: bool) . ((if w__2 then - (read_reg FPEXC_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__3 :: 32 Word.word) . + (read_reg FPEXC_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__3 :: 32 bits) . write_reg FPEXC_ref ((set_slice (( 32 :: int)::ii) (( 1 :: int)::ii) w__3 (( 29 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 32 Word.word))) @@ -21406,15 +21401,15 @@ definition aarch64_memory_single_general_register :: " AccType \<Rightarrow> in else (aget_X (( 8 :: int)::ii) t :: ( 8 Word.word) M)) \<bind> (\<lambda> (data :: 8 bits) . aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: + 8 Word.word) . (let data = w__4 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: ( 'regsize::len)Word.word) . aset_X t w__5) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: ( 'regsize::len)Word.word) . aset_X t w__6))) | MemOp_PREFETCH => @@ -21468,15 +21463,15 @@ definition aarch64_memory_single_general_register :: " AccType \<Rightarrow> in else (aget_X (( 16 :: int)::ii) t :: ( 16 Word.word) M)) \<bind> (\<lambda> (data :: 16 bits) . aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: + 16 Word.word) . (let data = w__12 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: ( 'regsize::len)Word.word) . aset_X t w__13) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14))) | MemOp_PREFETCH => @@ -21530,15 +21525,15 @@ definition aarch64_memory_single_general_register :: " AccType \<Rightarrow> in else (aget_X (( 32 :: int)::ii) t :: ( 32 Word.word) M)) \<bind> (\<lambda> (data :: 32 bits) . aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: + 32 Word.word) . (let data = w__20 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: ( 'regsize::len)Word.word) . aset_X t w__21) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: ( 'regsize::len)Word.word) . aset_X t w__22))) | MemOp_PREFETCH => @@ -21592,15 +21587,15 @@ definition aarch64_memory_single_general_register :: " AccType \<Rightarrow> in else (aget_X (( 64 :: int)::ii) t :: ( 64 Word.word) M)) \<bind> (\<lambda> (data :: 64 bits) . aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: + 64 Word.word) . (let data = w__28 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X t w__29) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: ( 'regsize::len)Word.word) . aset_X t w__30))) | MemOp_PREFETCH => @@ -21654,15 +21649,15 @@ definition aarch64_memory_single_general_register :: " AccType \<Rightarrow> in else (aget_X (( 128 :: int)::ii) t :: ( 128 Word.word) M)) \<bind> (\<lambda> (data :: 128 bits) . aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: + 128 Word.word) . (let data = w__36 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: ( 'regsize::len)Word.word) . aset_X t w__37) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: ( 'regsize::len)Word.word) . aset_X t w__38))) | MemOp_PREFETCH => @@ -21726,15 +21721,15 @@ definition aarch64_memory_single_general_immediate_unsigned :: " AccType \<Righ else (aget_X (( 8 :: int)::ii) t :: ( 8 Word.word) M)) \<bind> (\<lambda> (data :: 8 bits) . aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: + 8 Word.word) . (let data = w__4 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: ( 'regsize::len)Word.word) . aset_X t w__5) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: ( 'regsize::len)Word.word) . aset_X t w__6))) | MemOp_PREFETCH => @@ -21787,15 +21782,15 @@ definition aarch64_memory_single_general_immediate_unsigned :: " AccType \<Righ else (aget_X (( 16 :: int)::ii) t :: ( 16 Word.word) M)) \<bind> (\<lambda> (data :: 16 bits) . aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: + 16 Word.word) . (let data = w__12 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: ( 'regsize::len)Word.word) . aset_X t w__13) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14))) | MemOp_PREFETCH => @@ -21848,15 +21843,15 @@ definition aarch64_memory_single_general_immediate_unsigned :: " AccType \<Righ else (aget_X (( 32 :: int)::ii) t :: ( 32 Word.word) M)) \<bind> (\<lambda> (data :: 32 bits) . aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: + 32 Word.word) . (let data = w__20 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: ( 'regsize::len)Word.word) . aset_X t w__21) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: ( 'regsize::len)Word.word) . aset_X t w__22))) | MemOp_PREFETCH => @@ -21909,15 +21904,15 @@ definition aarch64_memory_single_general_immediate_unsigned :: " AccType \<Righ else (aget_X (( 64 :: int)::ii) t :: ( 64 Word.word) M)) \<bind> (\<lambda> (data :: 64 bits) . aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: + 64 Word.word) . (let data = w__28 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X t w__29) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: ( 'regsize::len)Word.word) . aset_X t w__30))) | MemOp_PREFETCH => @@ -21970,15 +21965,15 @@ definition aarch64_memory_single_general_immediate_unsigned :: " AccType \<Righ else (aget_X (( 128 :: int)::ii) t :: ( 128 Word.word) M)) \<bind> (\<lambda> (data :: 128 bits) . aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: + 128 Word.word) . (let data = w__36 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: ( 'regsize::len)Word.word) . aset_X t w__37) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: ( 'regsize::len)Word.word) . aset_X t w__38))) | MemOp_PREFETCH => @@ -22042,15 +22037,15 @@ definition aarch64_memory_single_general_immediate_signed_postidx :: " AccType else (aget_X (( 8 :: int)::ii) t :: ( 8 Word.word) M)) \<bind> (\<lambda> (data :: 8 bits) . aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: + 8 Word.word) . (let data = w__4 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: ( 'regsize::len)Word.word) . aset_X t w__5) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: ( 'regsize::len)Word.word) . aset_X t w__6))) | MemOp_PREFETCH => @@ -22103,15 +22098,15 @@ definition aarch64_memory_single_general_immediate_signed_postidx :: " AccType else (aget_X (( 16 :: int)::ii) t :: ( 16 Word.word) M)) \<bind> (\<lambda> (data :: 16 bits) . aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: + 16 Word.word) . (let data = w__12 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: ( 'regsize::len)Word.word) . aset_X t w__13) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14))) | MemOp_PREFETCH => @@ -22164,15 +22159,15 @@ definition aarch64_memory_single_general_immediate_signed_postidx :: " AccType else (aget_X (( 32 :: int)::ii) t :: ( 32 Word.word) M)) \<bind> (\<lambda> (data :: 32 bits) . aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: + 32 Word.word) . (let data = w__20 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: ( 'regsize::len)Word.word) . aset_X t w__21) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: ( 'regsize::len)Word.word) . aset_X t w__22))) | MemOp_PREFETCH => @@ -22225,15 +22220,15 @@ definition aarch64_memory_single_general_immediate_signed_postidx :: " AccType else (aget_X (( 64 :: int)::ii) t :: ( 64 Word.word) M)) \<bind> (\<lambda> (data :: 64 bits) . aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: + 64 Word.word) . (let data = w__28 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X t w__29) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: ( 'regsize::len)Word.word) . aset_X t w__30))) | MemOp_PREFETCH => @@ -22286,15 +22281,15 @@ definition aarch64_memory_single_general_immediate_signed_postidx :: " AccType else (aget_X (( 128 :: int)::ii) t :: ( 128 Word.word) M)) \<bind> (\<lambda> (data :: 128 bits) . aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: + 128 Word.word) . (let data = w__36 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: ( 'regsize::len)Word.word) . aset_X t w__37) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: ( 'regsize::len)Word.word) . aset_X t w__38))) | MemOp_PREFETCH => @@ -22350,7 +22345,7 @@ definition aarch64_memory_single_general_immediate_signed_pac :: " int \<Righta (aget_X (( 64 :: int)::ii) (( 31 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 Word.word) . (AuthDB address w__4 :: ( 64 Word.word) M))) \<bind> (\<lambda> (address :: 64 bits) . (let address = ((add_vec address offset :: 64 Word.word)) in - (aget_Mem address (( 8 :: int)::ii) AccType_NORMAL :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: 64 bits) . + (aget_Mem address (( 8 :: int)::ii) AccType_NORMAL :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: 64 Word.word) . (let data = w__6 in aset_X t data \<then> (if wback then @@ -22401,15 +22396,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_unpriv :: " Ac else (aget_X (( 8 :: int)::ii) t :: ( 8 Word.word) M)) \<bind> (\<lambda> (data :: 8 bits) . aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: + 8 Word.word) . (let data = w__4 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: ( 'regsize::len)Word.word) . aset_X t w__5) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: ( 'regsize::len)Word.word) . aset_X t w__6))) | MemOp_PREFETCH => @@ -22462,15 +22457,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_unpriv :: " Ac else (aget_X (( 16 :: int)::ii) t :: ( 16 Word.word) M)) \<bind> (\<lambda> (data :: 16 bits) . aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: + 16 Word.word) . (let data = w__12 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: ( 'regsize::len)Word.word) . aset_X t w__13) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14))) | MemOp_PREFETCH => @@ -22523,15 +22518,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_unpriv :: " Ac else (aget_X (( 32 :: int)::ii) t :: ( 32 Word.word) M)) \<bind> (\<lambda> (data :: 32 bits) . aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: + 32 Word.word) . (let data = w__20 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: ( 'regsize::len)Word.word) . aset_X t w__21) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: ( 'regsize::len)Word.word) . aset_X t w__22))) | MemOp_PREFETCH => @@ -22584,15 +22579,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_unpriv :: " Ac else (aget_X (( 64 :: int)::ii) t :: ( 64 Word.word) M)) \<bind> (\<lambda> (data :: 64 bits) . aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: + 64 Word.word) . (let data = w__28 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X t w__29) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: ( 'regsize::len)Word.word) . aset_X t w__30))) | MemOp_PREFETCH => @@ -22645,15 +22640,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_unpriv :: " Ac else (aget_X (( 128 :: int)::ii) t :: ( 128 Word.word) M)) \<bind> (\<lambda> (data :: 128 bits) . aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: + 128 Word.word) . (let data = w__36 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: ( 'regsize::len)Word.word) . aset_X t w__37) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: ( 'regsize::len)Word.word) . aset_X t w__38))) | MemOp_PREFETCH => @@ -22717,15 +22712,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_normal :: " Ac else (aget_X (( 8 :: int)::ii) t :: ( 8 Word.word) M)) \<bind> (\<lambda> (data :: 8 bits) . aset_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: 8 - bits) . + (aget_Mem address (((( 8 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__4 :: + 8 Word.word) . (let data = w__4 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: ( 'regsize::len)Word.word) . aset_X t w__5) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: ( 'regsize::len)Word.word) . aset_X t w__6))) | MemOp_PREFETCH => @@ -22778,15 +22773,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_normal :: " Ac else (aget_X (( 16 :: int)::ii) t :: ( 16 Word.word) M)) \<bind> (\<lambda> (data :: 16 bits) . aset_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: 16 - bits) . + (aget_Mem address (((( 16 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__12 :: + 16 Word.word) . (let data = w__12 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__13 :: ( 'regsize::len)Word.word) . aset_X t w__13) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14))) | MemOp_PREFETCH => @@ -22839,15 +22834,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_normal :: " Ac else (aget_X (( 32 :: int)::ii) t :: ( 32 Word.word) M)) \<bind> (\<lambda> (data :: 32 bits) . aset_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: 32 - bits) . + (aget_Mem address (((( 32 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__20 :: + 32 Word.word) . (let data = w__20 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__21 :: ( 'regsize::len)Word.word) . aset_X t w__21) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__22 :: ( 'regsize::len)Word.word) . aset_X t w__22))) | MemOp_PREFETCH => @@ -22900,15 +22895,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_normal :: " Ac else (aget_X (( 64 :: int)::ii) t :: ( 64 Word.word) M)) \<bind> (\<lambda> (data :: 64 bits) . aset_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: 64 - bits) . + (aget_Mem address (((( 64 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: + 64 Word.word) . (let data = w__28 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X t w__29) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__30 :: ( 'regsize::len)Word.word) . aset_X t w__30))) | MemOp_PREFETCH => @@ -22961,15 +22956,15 @@ definition aarch64_memory_single_general_immediate_signed_offset_normal :: " Ac else (aget_X (( 128 :: int)::ii) t :: ( 128 Word.word) M)) \<bind> (\<lambda> (data :: 128 bits) . aset_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype data) | MemOp_LOAD => - (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: 128 - bits) . + (aget_Mem address (((( 128 :: int)::ii) div (( 8 :: int)::ii))) acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__36 :: + 128 Word.word) . (let data = w__36 in if signed then - (SignExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: + (SignExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__37 :: ( 'regsize::len)Word.word) . aset_X t w__37) else - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__38 :: ( 'regsize::len)Word.word) . aset_X t w__38))) | MemOp_PREFETCH => @@ -23003,8 +22998,8 @@ definition aarch64_memory_pair_simdfp_postidx :: " AccType \<Rightarrow>('datas assert_exp True (''dbytes constraint'')) \<then> CheckFPAdvSIMDEnabled64 () ) \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data1 :: 'datasize bits) . - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data2 :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data1 . + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data2 . (let (rt_unknown :: bool) = False in (if ((((((memop = MemOp_LOAD))) \<and> (((t = t2)))))) then (let (c :: Constraint) = (ConstrainUnpredictable Unpredictable_LDPOVERLAP) in @@ -23022,29 +23017,28 @@ definition aarch64_memory_pair_simdfp_postidx :: " AccType \<Rightarrow>('datas else address) in (case memop of MemOp_STORE => - (aget_V ((int (size data1))) t :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'datasize bits) . + (aget_V ((int (size data1))) t :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: ( 'datasize::len)Word.word) . (let data1 = w__2 in - (aget_V ((int (size data1))) t2 :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: 'datasize bits) . + (aget_V ((int (size data1))) t2 :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'datasize::len)Word.word) . (let data2 = w__3 in aset_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype data1 \<then> aset_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype data2)))) | MemOp_LOAD => - (aget_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: 'datasize - bits) . + (aget_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: + ( 'datasize::len)Word.word) . (let data1 = w__4 in - (aget_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: 'datasize - bits) . + (aget_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + ( 'datasize::len)Word.word) . (let data2 = w__5 in (if rt_unknown then - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: 'datasize - bits) . + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + ( 'datasize::len)Word.word) . (let data1 = w__6 in - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: 'datasize - bits) . - (let (data2 :: 'datasize bits) = w__7 in + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: + ( 'datasize::len)Word.word) . + (let data2 = w__7 in return (data1, data2))))) - else return (data1, data2)) \<bind> (\<lambda> varstup . (let ((data1 :: 'datasize bits), (data2 :: 'datasize - bits)) = varstup in + else return (data1, data2)) \<bind> (\<lambda> varstup . (let (data1, data2) = varstup in aset_V t data1 \<then> aset_V t2 data2)))))) ) \<then> (if wback then @@ -23066,8 +23060,8 @@ definition aarch64_memory_pair_simdfp_noalloc :: " AccType \<Rightarrow>('datas assert_exp True (''dbytes constraint'')) \<then> CheckFPAdvSIMDEnabled64 () ) \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data1 :: 'datasize bits) . - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data2 :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data1 . + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data2 . (let (rt_unknown :: bool) = False in (if ((((((memop = MemOp_LOAD))) \<and> (((t = t2)))))) then (let (c :: Constraint) = (ConstrainUnpredictable Unpredictable_LDPOVERLAP) in @@ -23085,29 +23079,28 @@ definition aarch64_memory_pair_simdfp_noalloc :: " AccType \<Rightarrow>('datas else address) in (case memop of MemOp_STORE => - (aget_V ((int (size data1))) t :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: 'datasize bits) . + (aget_V ((int (size data1))) t :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__2 :: ( 'datasize::len)Word.word) . (let data1 = w__2 in - (aget_V ((int (size data1))) t2 :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: 'datasize bits) . + (aget_V ((int (size data1))) t2 :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'datasize::len)Word.word) . (let data2 = w__3 in aset_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype data1 \<then> aset_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype data2)))) | MemOp_LOAD => - (aget_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: 'datasize - bits) . + (aget_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__4 :: + ( 'datasize::len)Word.word) . (let data1 = w__4 in - (aget_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: 'datasize - bits) . + (aget_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__5 :: + ( 'datasize::len)Word.word) . (let data2 = w__5 in (if rt_unknown then - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: 'datasize - bits) . + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + ( 'datasize::len)Word.word) . (let data1 = w__6 in - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: 'datasize - bits) . - (let (data2 :: 'datasize bits) = w__7 in + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: + ( 'datasize::len)Word.word) . + (let data2 = w__7 in return (data1, data2))))) - else return (data1, data2)) \<bind> (\<lambda> varstup . (let ((data1 :: 'datasize bits), (data2 :: 'datasize - bits)) = varstup in + else return (data1, data2)) \<bind> (\<lambda> varstup . (let (data1, data2) = varstup in aset_V t data1 \<then> aset_V t2 data2)))))) ) \<then> (if wback then @@ -23129,8 +23122,8 @@ definition aarch64_memory_pair_general_postidx :: " AccType \<Rightarrow>('data assert_exp True (''dbytes constraint'')) \<then> ((let wback = wback__arg in (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (address :: 64 bits) . - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data1 :: 'datasize bits) . - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data2 :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data1 . + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data2 . (let (rt_unknown :: bool) = False in (let (wb_unknown :: bool) = False in (if ((((((((((((memop = MemOp_LOAD))) \<and> wback))) \<and> ((((((t = n))) \<or> (((t2 = n))))))))) \<and> (((n \<noteq> (( 31 :: int)::ii))))))) then @@ -23175,29 +23168,28 @@ definition aarch64_memory_pair_general_postidx :: " AccType \<Rightarrow>('data MemOp_STORE => (if (((rt_unknown \<and> (((t = n)))))) then (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) - else (aget_X ((int (size data1))) t :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> (data1 :: 'datasize bits) . + else (aget_X ((int (size data1))) t :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> data1 . (if (((rt_unknown \<and> (((t2 = n)))))) then (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) - else (aget_X ((int (size data1))) t2 :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> (data2 :: 'datasize bits) . + else (aget_X ((int (size data1))) t2 :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> data2 . aset_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype data1 \<then> aset_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype data2)) | MemOp_LOAD => - (aget_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: 'datasize - bits) . + (aget_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + ( 'datasize::len)Word.word) . (let data1 = w__6 in - (aget_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: 'datasize - bits) . + (aget_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: + ( 'datasize::len)Word.word) . (let data2 = w__7 in (if rt_unknown then - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__8 :: 'datasize - bits) . + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__8 :: + ( 'datasize::len)Word.word) . (let data1 = w__8 in - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__9 :: 'datasize - bits) . - (let (data2 :: 'datasize bits) = w__9 in + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__9 :: + ( 'datasize::len)Word.word) . + (let data2 = w__9 in return (data1, data2))))) - else return (data1, data2)) \<bind> (\<lambda> varstup . (let ((data1 :: 'datasize bits), (data2 :: 'datasize - bits)) = varstup in + else return (data1, data2)) \<bind> (\<lambda> varstup . (let (data1, data2) = varstup in if signed then (SignExtend__0 data1 ((make_the_value (( 64 :: int)::ii) :: 64 itself)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__10 :: 64 Word.word) . @@ -23228,8 +23220,8 @@ definition aarch64_memory_pair_general_noalloc :: " AccType \<Rightarrow>('data ((assert_exp True (''datasize constraint'') \<then> assert_exp True (''dbytes constraint'')) \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (address :: 64 bits) . - (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data1 :: 'datasize bits) . - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (data2 :: 'datasize bits) . + (undefined_bitvector datasize :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data1 . + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> data2 . (let (rt_unknown :: bool) = False in (if ((((((memop = MemOp_LOAD))) \<and> (((t = t2)))))) then (let (c :: Constraint) = (ConstrainUnpredictable Unpredictable_LDPOVERLAP) in @@ -23249,29 +23241,28 @@ definition aarch64_memory_pair_general_noalloc :: " AccType \<Rightarrow>('data MemOp_STORE => (if (((rt_unknown \<and> (((t = n)))))) then (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) - else (aget_X ((int (size data1))) t :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> (data1 :: 'datasize bits) . + else (aget_X ((int (size data1))) t :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> data1 . (if (((rt_unknown \<and> (((t2 = n)))))) then (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) - else (aget_X ((int (size data1))) t2 :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> (data2 :: 'datasize bits) . + else (aget_X ((int (size data1))) t2 :: (( 'datasize::len)Word.word) M)) \<bind> (\<lambda> data2 . aset_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype data1 \<then> aset_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype data2)) | MemOp_LOAD => - (aget_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: 'datasize - bits) . + (aget_Mem ((add_vec_int address (( 0 :: int)::ii) :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__6 :: + ( 'datasize::len)Word.word) . (let data1 = w__6 in - (aget_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: 'datasize - bits) . + (aget_Mem ((add_vec_int address dbytes :: 64 Word.word)) dbytes acctype :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__7 :: + ( 'datasize::len)Word.word) . (let data2 = w__7 in (if rt_unknown then - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__8 :: 'datasize - bits) . + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__8 :: + ( 'datasize::len)Word.word) . (let data1 = w__8 in - (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__9 :: 'datasize - bits) . - (let (data2 :: 'datasize bits) = w__9 in + (undefined_bitvector ((int (size data1))) :: (( 'datasize::len)Word.word) M) \<bind> (\<lambda> (w__9 :: + ( 'datasize::len)Word.word) . + (let data2 = w__9 in return (data1, data2))))) - else return (data1, data2)) \<bind> (\<lambda> varstup . (let ((data1 :: 'datasize bits), (data2 :: 'datasize - bits)) = varstup in + else return (data1, data2)) \<bind> (\<lambda> varstup . (let (data1, data2) = varstup in aset_X t data1 \<then> aset_X t2 data2)))))) ) \<then> (if wback then @@ -23348,7 +23339,7 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize (if pair then assert_exp True ('''') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 Word.word) . aset_X t w__9) else ((if (((address \<noteq> ((Align__1 address dbytes :: 64 Word.word))))) then @@ -23364,9 +23355,9 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize 64 Word.word) . aset_X t2 w__12))) else - (aget_Mem address dbytes acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__13 :: 8 bits) . + (aget_Mem address dbytes acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__13 :: 8 Word.word) . (let data = w__13 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14)))) ))))))))))))) @@ -23431,7 +23422,7 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize (if pair then assert_exp True ('''') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__24 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__24 :: 32 Word.word) . aset_X t w__24) else ((if (((address \<noteq> ((Align__1 address dbytes :: 64 Word.word))))) then @@ -23447,9 +23438,9 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize 64 Word.word) . aset_X t2 w__27))) else - (aget_Mem address dbytes acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__28 :: 16 bits) . + (aget_Mem address dbytes acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__28 :: 16 Word.word) . (let data = w__28 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X t w__29)))) ))))))))))))) @@ -23514,7 +23505,7 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize (if pair then assert_exp True ('''') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__39 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__39 :: 32 Word.word) . aset_X t w__39) else ((if (((address \<noteq> ((Align__1 address dbytes :: 64 Word.word))))) then @@ -23530,9 +23521,9 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize 64 Word.word) . aset_X t2 w__42))) else - (aget_Mem address dbytes acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__43 :: 32 bits) . + (aget_Mem address dbytes acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__43 :: 32 Word.word) . (let data = w__43 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__44 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__44 :: ( 'regsize::len)Word.word) . aset_X t w__44)))) ))))))))))))) @@ -23597,10 +23588,10 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize (if pair then assert_exp True ('''') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__54 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__54 :: 32 Word.word) . aset_X t w__54) else if (((elsize = (( 32 :: int)::ii)))) then - (aget_Mem address dbytes acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__55 :: 64 bits) . + (aget_Mem address dbytes acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__55 :: 64 Word.word) . (let data = w__55 in BigEndian () \<bind> (\<lambda> (w__56 :: bool) . if w__56 then @@ -23625,9 +23616,9 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize 64 Word.word) . aset_X t2 w__59))) else - (aget_Mem address dbytes acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__60 :: 64 bits) . + (aget_Mem address dbytes acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__60 :: 64 Word.word) . (let data = w__60 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__61 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__61 :: ( 'regsize::len)Word.word) . aset_X t w__61)))) ))))))))))))) @@ -23692,10 +23683,10 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize (if pair then assert_exp True ('''') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__71 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__71 :: 32 Word.word) . aset_X t w__71) else if (((elsize = (( 32 :: int)::ii)))) then - (aget_Mem address dbytes acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__72 :: 128 bits) . + (aget_Mem address dbytes acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__72 :: 128 Word.word) . (let data = w__72 in BigEndian () \<bind> (\<lambda> (w__73 :: bool) . if w__73 then @@ -23720,9 +23711,9 @@ definition aarch64_memory_exclusive_single :: " AccType \<Rightarrow>('datasize 64 Word.word) . aset_X t2 w__76))) else - (aget_Mem address dbytes acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__77 :: 128 bits) . + (aget_Mem address dbytes acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__77 :: 128 Word.word) . (let data = w__77 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__78 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__78 :: ( 'regsize::len)Word.word) . aset_X t w__78)))) ))))))))))))) @@ -23877,7 +23868,7 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: (if pair then assert_exp True (''datasize constraint'') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__9 :: 32 Word.word) . aset_X t w__9) else ((if (((address \<noteq> ((Align__1 address dbytes :: 64 Word.word))))) then @@ -23893,9 +23884,9 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: 64 Word.word) . aset_X t2 w__12))) else - (aget_Mem address dbytes acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__13 :: 8 bits) . + (aget_Mem address dbytes acctype :: ( 8 Word.word) M) \<bind> (\<lambda> (w__13 :: 8 Word.word) . (let data = w__13 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__14 :: ( 'regsize::len)Word.word) . aset_X t w__14)))) ))))))))))))) @@ -23960,7 +23951,7 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: (if pair then assert_exp True (''datasize constraint'') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__24 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__24 :: 32 Word.word) . aset_X t w__24) else ((if (((address \<noteq> ((Align__1 address dbytes :: 64 Word.word))))) then @@ -23976,9 +23967,9 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: 64 Word.word) . aset_X t2 w__27))) else - (aget_Mem address dbytes acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__28 :: 16 bits) . + (aget_Mem address dbytes acctype :: ( 16 Word.word) M) \<bind> (\<lambda> (w__28 :: 16 Word.word) . (let data = w__28 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__29 :: ( 'regsize::len)Word.word) . aset_X t w__29)))) ))))))))))))) @@ -24043,7 +24034,7 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: (if pair then assert_exp True (''datasize constraint'') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__39 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__39 :: 32 Word.word) . aset_X t w__39) else ((if (((address \<noteq> ((Align__1 address dbytes :: 64 Word.word))))) then @@ -24059,9 +24050,9 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: 64 Word.word) . aset_X t2 w__42))) else - (aget_Mem address dbytes acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__43 :: 32 bits) . + (aget_Mem address dbytes acctype :: ( 32 Word.word) M) \<bind> (\<lambda> (w__43 :: 32 Word.word) . (let data = w__43 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__44 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__44 :: ( 'regsize::len)Word.word) . aset_X t w__44)))) ))))))))))))) @@ -24126,10 +24117,10 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: (if pair then assert_exp True (''datasize constraint'') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__54 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__54 :: 32 Word.word) . aset_X t w__54) else if (((elsize = (( 32 :: int)::ii)))) then - (aget_Mem address dbytes acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__55 :: 64 bits) . + (aget_Mem address dbytes acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__55 :: 64 Word.word) . (let data = w__55 in BigEndian () \<bind> (\<lambda> (w__56 :: bool) . if w__56 then @@ -24154,9 +24145,9 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: 64 Word.word) . aset_X t2 w__59))) else - (aget_Mem address dbytes acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__60 :: 64 bits) . + (aget_Mem address dbytes acctype :: ( 64 Word.word) M) \<bind> (\<lambda> (w__60 :: 64 Word.word) . (let data = w__60 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__61 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__61 :: ( 'regsize::len)Word.word) . aset_X t w__61)))) ))))))))))))) @@ -24221,10 +24212,10 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: (if pair then assert_exp True (''datasize constraint'') \<then> (if rt_unknown then - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__71 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__71 :: 32 Word.word) . aset_X t w__71) else if (((elsize = (( 32 :: int)::ii)))) then - (aget_Mem address dbytes acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__72 :: 128 bits) . + (aget_Mem address dbytes acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__72 :: 128 Word.word) . (let data = w__72 in BigEndian () \<bind> (\<lambda> (w__73 :: bool) . if w__73 then @@ -24249,9 +24240,9 @@ definition aarch64_memory_exclusive_pair :: " AccType \<Rightarrow>('datasize:: 64 Word.word) . aset_X t2 w__76))) else - (aget_Mem address dbytes acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__77 :: 128 bits) . + (aget_Mem address dbytes acctype :: ( 128 Word.word) M) \<bind> (\<lambda> (w__77 :: 128 Word.word) . (let data = w__77 in - (ZeroExtend__0 data ((make_the_value regsize :: ( 'regsize::len)itself)) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__78 :: + (ZeroExtend__0 data ((make_the_value regsize )) :: (( 'regsize::len)Word.word) M) \<bind> (\<lambda> (w__78 :: ( 'regsize::len)Word.word) . aset_X t w__78)))) ))))))))))))) @@ -30572,7 +30563,7 @@ definition memory_literal_simdfp_decode :: "(2)Word.word \<Rightarrow>(1)Word.w else UnallocatedEncoding () \<then> return size1) \<bind> (\<lambda> (size1 :: ii) . (SignExtend__0 ((concat_vec imm19 (vec_of_bits [B0,B0] :: 2 Word.word) :: 21 Word.word)) ((make_the_value (( 64 :: int)::ii) :: 64 itself)) - :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . + :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . (let offset = w__0 in aarch64_memory_literal_simdfp offset ((ex_int size1)) t)))))))))" @@ -31348,7 +31339,7 @@ definition float_convert_int_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (fltsize, op1, part, rounding, unsigned)))) else if (((v__98 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__0) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31611,7 +31602,7 @@ definition float_convert_fix_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (op1, rounding, unsigned)))) else if (((b__2 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__0) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31644,7 +31635,7 @@ definition float_convert_fix_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (op1, rounding, unsigned)))) else if (((b__6 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__1) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31677,7 +31668,7 @@ definition float_convert_fix_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (op1, rounding, unsigned)))) else if (((b__10 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__2 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__2) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31710,7 +31701,7 @@ definition float_convert_fix_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (op1, rounding, unsigned)))) else if (((b__14 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__3 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__3 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__3) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31743,7 +31734,7 @@ definition float_convert_fix_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (op1, rounding, unsigned)))) else if (((b__18 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__4 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__4) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31776,7 +31767,7 @@ definition float_convert_fix_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (op1, rounding, unsigned)))) else if (((b__22 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__5 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__5) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31809,7 +31800,7 @@ definition float_convert_fix_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (op1, rounding, unsigned)))) else if (((b__26 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__6) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31842,7 +31833,7 @@ definition float_convert_fix_decode :: "(1)Word.word \<Rightarrow>(1)Word.word (let (op1 :: FPConvOp) = FPConvOp_CVT_FtoI in return (op1, rounding, unsigned)))) else if (((b__30 = (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__7 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__7 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__7) in (let (unsigned :: bool) = ((vec_of_bits [access_vec_dec opcode (( 0 :: int)::ii)] :: 1 Word.word) = (vec_of_bits [B1] :: 1 Word.word)) in @@ -31950,12 +31941,12 @@ definition float_arithmetic_round_decode :: "(1)Word.word \<Rightarrow>(1)Word. else if (((v__101 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then UnallocatedEncoding () \<then> return (exact, rounding) else if (((v__101 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__0) in (let (exact :: bool) = True in return (exact, rounding)))) else - (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . + (read_reg FPCR_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . (let (rounding :: FPRounding) = (FPRoundingMode w__1) in return (exact, rounding)))) \<bind> (\<lambda> varstup . (let ((exact :: bool), (rounding :: FPRounding)) = varstup in aarch64_float_arithmetic_round d datasize exact n rounding))))))))))))" @@ -33643,7 +33634,7 @@ definition integer_logical_immediate_decode :: "(1)Word.word \<Rightarrow>(2)Wo else return () ) \<then> (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (anon10 :: 32 bits) . (DecodeBitMasks (( 32 :: int)::ii) N imms immr True :: (( 32 Word.word * 32 Word.word)) M) \<bind> (\<lambda> (w__0 :: - ( 32 bits * 32 bits)) . + ( 32 Word.word * 32 Word.word)) . (let (tup__0, tup__1) = w__0 in (let imm = tup__0 in (let anon10 = tup__1 in @@ -33683,7 +33674,7 @@ definition integer_logical_immediate_decode :: "(1)Word.word \<Rightarrow>(2)Wo else return () ) \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (anon10 :: 64 bits) . (DecodeBitMasks (( 64 :: int)::ii) N imms immr True :: (( 64 Word.word * 64 Word.word)) M) \<bind> (\<lambda> (w__1 :: - ( 64 bits * 64 bits)) . + ( 64 Word.word * 64 Word.word)) . (let (tup__0, tup__1) = w__1 in (let imm = tup__0 in (let anon10 = tup__1 in @@ -33727,7 +33718,7 @@ definition integer_bitfield_decode :: "(1)Word.word \<Rightarrow>(2)Word.word \ ((let R1 = (Word.uint immr) in (let S = (Word.uint imms) in (DecodeBitMasks (( 32 :: int)::ii) N imms immr False :: (( 32 Word.word * 32 Word.word)) M) \<bind> (\<lambda> (w__0 :: - ( 32 bits * 32 bits)) . + ( 32 Word.word * 32 Word.word)) . (let (tup__0, tup__1) = w__0 in (let wmask = tup__0 in (let tmask = tup__1 in @@ -33766,7 +33757,7 @@ definition integer_bitfield_decode :: "(1)Word.word \<Rightarrow>(2)Word.word \ ((let R1 = (Word.uint immr) in (let S = (Word.uint imms) in (DecodeBitMasks (( 64 :: int)::ii) N imms immr False :: (( 64 Word.word * 64 Word.word)) M) \<bind> (\<lambda> (w__1 :: - ( 64 bits * 64 bits)) . + ( 64 Word.word * 64 Word.word)) . (let (tup__0, tup__1) = w__1 in (let wmask = tup__0 in (let tmask = tup__1 in @@ -34746,17 +34737,67 @@ definition decode :: "(32)Word.word \<Rightarrow>((register_value),(unit),(exce integer_arithmetic_div_decode sf op1 S Rm opcode2 o1 Rn Rd)))))))))" +(*val fetch_and_execute : unit -> M unit*) + +definition fetch_and_execute :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where + " fetch_and_execute _ = ( + (whileM () + (\<lambda> unit_var . return True) + (\<lambda> unit_var . + (try_catch ((read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . + (aget_Mem w__0 (( 4 :: int)::ii) AccType_IFETCH :: ( 32 Word.word) M) \<bind> (\<lambda> instr . + decode instr))) (\<lambda>x . + (case x of + Error_Undefined _ => exit0 () + | Error_See s => + if(s = (''HINT'')) then (return () ) else (exit0 () ) + | Error_Implementation_Defined _ => exit0 () + | Error_ReservedEncoding _ => exit0 () + | Error_ExceptionTaken _ => exit0 () + )) \<then> + read_reg BranchTaken_ref) \<bind> (\<lambda> (w__1 :: bool) . + if w__1 then write_reg BranchTaken_ref False + else + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . + write_reg PC_ref ((add_vec_int w__2 (( 4 :: int)::ii) :: 64 Word.word)))))))" + + +(*val main : unit -> M unit*) + +definition main :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where + " main _ = ( + (write_reg + PC_ref + ((GetSlice_int ((make_the_value (( 64 :: int)::ii) :: 64 itself)) ((elf_entry () )) (( 0 :: int)::ii) + :: 64 Word.word)) \<then> + (ZeroExtend__0 (vec_of_bits [B0,B0,B1,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 16 Word.word) + ((make_the_value (( 64 :: int)::ii) :: 64 itself)) + :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (write_reg SP_EL0_ref w__0 \<then> + read_reg PSTATE_ref) \<bind> (\<lambda> (w__1 :: ProcState) . + (write_reg PSTATE_ref (w__1 (| ProcState_D := ((vec_of_bits [B1] :: 1 Word.word))|)) \<then> + read_reg PSTATE_ref) \<bind> (\<lambda> (w__2 :: ProcState) . + (write_reg PSTATE_ref (w__2 (| ProcState_A := ((vec_of_bits [B1] :: 1 Word.word))|)) \<then> + read_reg PSTATE_ref) \<bind> (\<lambda> (w__3 :: ProcState) . + (write_reg PSTATE_ref (w__3 (| ProcState_I := ((vec_of_bits [B1] :: 1 Word.word))|)) \<then> + read_reg PSTATE_ref) \<bind> (\<lambda> (w__4 :: ProcState) . + (write_reg PSTATE_ref (w__4 (| ProcState_F := ((vec_of_bits [B1] :: 1 Word.word))|)) \<then> + (ZeroExtend__0 (vec_of_bits [B1,B0] :: 2 Word.word) ((make_the_value (( 32 :: int)::ii) :: 32 itself)) + :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__5 :: 32 Word.word) . + (write_reg OSLSR_EL1_ref w__5 \<then> write_reg BranchTaken_ref False) \<then> fetch_and_execute () )))))))" + + (*val initialize_registers : unit -> M unit*) definition initialize_registers :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " initialize_registers _ = ( undefined_bool () \<bind> (\<lambda> (w__0 :: bool) . (write_reg unconditional_ref w__0 \<then> - (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M)) \<bind> (\<lambda> (w__1 :: 4 bits) . + (undefined_bitvector (( 4 :: int)::ii) :: ( 4 Word.word) M)) \<bind> (\<lambda> (w__1 :: 4 Word.word) . (write_reg currentCond_ref w__1 \<then> undefined___InstrEnc () ) \<bind> (\<lambda> (w__2 :: InstrEnc) . (write_reg ThisInstrEnc_ref w__2 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__3 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__3 :: 32 Word.word) . (write_reg ThisInstr_ref w__3 \<then> undefined_bool () ) \<bind> (\<lambda> (w__4 :: bool) . (write_reg Sleeping_ref w__4 \<then> @@ -34764,246 +34805,246 @@ definition initialize_registers :: " unit \<Rightarrow>((register_value),(unit) (write_reg PendingPhysicalSError_ref w__5 \<then> undefined_bool () ) \<bind> (\<lambda> (w__6 :: bool) . (write_reg PendingInterrupt_ref w__6 \<then> - (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M)) \<bind> (\<lambda> (w__7 :: 52 bits) . + (undefined_bitvector (( 52 :: int)::ii) :: ( 52 Word.word) M)) \<bind> (\<lambda> (w__7 :: 52 Word.word) . (write_reg Memory_ref w__7 \<then> undefined_bool () ) \<bind> (\<lambda> (w__8 :: bool) . (write_reg ExclusiveLocal_ref w__8 \<then> undefined_bool () ) \<bind> (\<lambda> (w__9 :: bool) . (write_reg BranchTaken_ref w__9 \<then> (undefined_bitvector (( 128 :: int)::ii) :: ( 128 Word.word) M)) \<bind> (\<lambda> (w__10 :: 128 Word.word) . - (undefined_vector (( 32 :: int)::ii) w__10 :: ( ( 128 Word.word)list) M) \<bind> (\<lambda> (w__11 :: ( 128 bits) list) . + (undefined_vector (( 32 :: int)::ii) w__10 :: ( ( 128 Word.word)list) M) \<bind> (\<lambda> (w__11 :: ( 128 Word.word) list) . (write_reg V_ref w__11 \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__12 :: 64 Word.word) . - (undefined_vector (( 31 :: int)::ii) w__12 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__13 :: ( 64 bits) list) . + (undefined_vector (( 31 :: int)::ii) w__12 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__13 :: ( 64 Word.word) list) . (write_reg R_ref w__13 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__14 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__14 :: 64 Word.word) . (write_reg PC_ref w__14 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__15 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__15 :: 64 Word.word) . (write_reg VTTBR_EL2_ref w__15 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__16 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__16 :: 32 Word.word) . (write_reg VTCR_EL2_ref w__16 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__17 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__17 :: 32 Word.word) . (write_reg VSESR_EL2_ref w__17 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__18 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__18 :: 32 Word.word) . (write_reg VDFSR_ref w__18 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__19 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__19 :: 64 Word.word) . (write_reg VBAR_EL3_ref w__19 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__20 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__20 :: 64 Word.word) . (write_reg VBAR_EL2_ref w__20 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__21 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__21 :: 64 Word.word) . (write_reg VBAR_EL1_ref w__21 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__22 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__22 :: 32 Word.word) . (write_reg VBAR_ref w__22 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__23 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__23 :: 64 Word.word) . (write_reg TTBR1_EL2_ref w__23 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__24 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__24 :: 64 Word.word) . (write_reg TTBR1_EL1_ref w__24 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__25 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__25 :: 64 Word.word) . (write_reg TTBR0_EL3_ref w__25 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__26 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__26 :: 64 Word.word) . (write_reg TTBR0_EL2_ref w__26 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__27 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__27 :: 64 Word.word) . (write_reg TTBR0_EL1_ref w__27 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__28 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__28 :: 32 Word.word) . (write_reg TTBCR_ref w__28 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__29 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__29 :: 32 Word.word) . (write_reg TCR_EL3_ref w__29 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__30 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__30 :: 64 Word.word) . (write_reg TCR_EL2_ref w__30 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__31 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__31 :: 64 Word.word) . (write_reg TCR_EL1_ref w__31 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__32 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__32 :: 32 Word.word) . (write_reg SP_mon_ref w__32 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__33 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__33 :: 64 Word.word) . (write_reg SP_EL3_ref w__33 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__34 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__34 :: 64 Word.word) . (write_reg SP_EL2_ref w__34 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__35 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__35 :: 64 Word.word) . (write_reg SP_EL1_ref w__35 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__36 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__36 :: 64 Word.word) . (write_reg SP_EL0_ref w__36 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__37 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__37 :: 32 Word.word) . (write_reg SPSR_und_ref w__37 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__38 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__38 :: 32 Word.word) . (write_reg SPSR_svc_ref w__38 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__39 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__39 :: 32 Word.word) . (write_reg SPSR_mon_ref w__39 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__40 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__40 :: 32 Word.word) . (write_reg SPSR_irq_ref w__40 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__41 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__41 :: 32 Word.word) . (write_reg SPSR_hyp_ref w__41 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__42 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__42 :: 32 Word.word) . (write_reg SPSR_fiq_ref w__42 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__43 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__43 :: 32 Word.word) . (write_reg SPSR_abt_ref w__43 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__44 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__44 :: 32 Word.word) . (write_reg SPSR_EL3_ref w__44 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__45 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__45 :: 32 Word.word) . (write_reg SPSR_EL2_ref w__45 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__46 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__46 :: 32 Word.word) . (write_reg SPSR_EL1_ref w__46 \<then> undefined_signal () ) \<bind> (\<lambda> (w__47 :: signal) . (write_reg SPIDEN_ref w__47 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__48 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__48 :: 32 Word.word) . (write_reg SDER_ref w__48 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__49 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__49 :: 32 Word.word) . (write_reg SDCR_ref w__49 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__50 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__50 :: 32 Word.word) . (write_reg SCTLR_EL3_ref w__50 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__51 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__51 :: 32 Word.word) . (write_reg SCTLR_EL2_ref w__51 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__52 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__52 :: 32 Word.word) . (write_reg SCTLR_EL1_ref w__52 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__53 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__53 :: 32 Word.word) . (write_reg SCTLR_ref w__53 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__54 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__54 :: 32 Word.word) . (write_reg SCR_EL3_ref w__54 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__55 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__55 :: 32 Word.word) . (write_reg SCR_ref w__55 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__56 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__56 :: 64 Word.word) . (write_reg RVBAR_EL3_ref w__56 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__57 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__57 :: 64 Word.word) . (write_reg RVBAR_EL2_ref w__57 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__58 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__58 :: 64 Word.word) . (write_reg RVBAR_EL1_ref w__58 \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__59 :: 64 Word.word) . - (undefined_vector (( 5 :: int)::ii) w__59 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__60 :: ( 64 bits) list) . + (undefined_vector (( 5 :: int)::ii) w__59 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__60 :: ( 64 Word.word) list) . (write_reg RC_ref w__60 \<then> undefined_ProcState () ) \<bind> (\<lambda> (w__61 :: ProcState) . (write_reg PSTATE_ref w__61 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__62 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__62 :: 32 Word.word) . (write_reg OSLSR_EL1_ref w__62 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__63 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__63 :: 32 Word.word) . (write_reg OSDLR_EL1_ref w__63 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__64 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__64 :: 32 Word.word) . (write_reg MDSCR_EL1_ref w__64 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__65 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__65 :: 32 Word.word) . (write_reg MDCR_EL3_ref w__65 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__66 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__66 :: 32 Word.word) . (write_reg MDCR_EL2_ref w__66 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__67 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__67 :: 64 Word.word) . (write_reg MAIR_EL3_ref w__67 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__68 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__68 :: 64 Word.word) . (write_reg MAIR_EL2_ref w__68 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__69 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__69 :: 64 Word.word) . (write_reg MAIR_EL1_ref w__69 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__70 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__70 :: 32 Word.word) . (write_reg LR_mon_ref w__70 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__71 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__71 :: 64 Word.word) . (write_reg ID_AA64DFR0_EL1_ref w__71 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__72 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__72 :: 32 Word.word) . (write_reg HVBAR_ref w__72 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__73 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__73 :: 32 Word.word) . (write_reg HSR_ref w__73 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__74 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__74 :: 32 Word.word) . (write_reg HSCTLR_ref w__74 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__75 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__75 :: 64 Word.word) . (write_reg HPFAR_EL2_ref w__75 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__76 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__76 :: 32 Word.word) . (write_reg HPFAR_ref w__76 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__77 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__77 :: 32 Word.word) . (write_reg HIFAR_ref w__77 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__78 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__78 :: 32 Word.word) . (write_reg HDFAR_ref w__78 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__79 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__79 :: 32 Word.word) . (write_reg HDCR_ref w__79 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__80 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__80 :: 64 Word.word) . (write_reg HCR_EL2_ref w__80 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__81 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__81 :: 32 Word.word) . (write_reg HCR2_ref w__81 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__82 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__82 :: 32 Word.word) . (write_reg HCR_ref w__82 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__83 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__83 :: 32 Word.word) . (write_reg FPSR_ref w__83 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__84 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__84 :: 32 Word.word) . (write_reg FPSCR_ref w__84 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__85 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__85 :: 32 Word.word) . (write_reg FPEXC_ref w__85 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__86 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__86 :: 32 Word.word) . (write_reg FPCR_ref w__86 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__87 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__87 :: 64 Word.word) . (write_reg FAR_EL3_ref w__87 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__88 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__88 :: 64 Word.word) . (write_reg FAR_EL2_ref w__88 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__89 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__89 :: 64 Word.word) . (write_reg FAR_EL1_ref w__89 \<then> - (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__90 :: 1 bits) . + (undefined_bitvector (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__90 :: 1 Word.word) . (write_reg EventRegister_ref w__90 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__91 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__91 :: 32 Word.word) . (write_reg ESR_EL3_ref w__91 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__92 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__92 :: 32 Word.word) . (write_reg ESR_EL2_ref w__92 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__93 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__93 :: 32 Word.word) . (write_reg ESR_EL1_ref w__93 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__94 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__94 :: 32 Word.word) . (write_reg ELR_hyp_ref w__94 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__95 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__95 :: 64 Word.word) . (write_reg ELR_EL3_ref w__95 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__96 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__96 :: 64 Word.word) . (write_reg ELR_EL2_ref w__96 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__97 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__97 :: 64 Word.word) . (write_reg ELR_EL1_ref w__97 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__98 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__98 :: 32 Word.word) . (write_reg EDSCR_ref w__98 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__99 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__99 :: 32 Word.word) . (write_reg DSPSR_EL0_ref w__99 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__100 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__100 :: 32 Word.word) . (write_reg DSPSR_ref w__100 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__101 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__101 :: 64 Word.word) . (write_reg DLR_EL0_ref w__101 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__102 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__102 :: 32 Word.word) . (write_reg DLR_ref w__102 \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__103 :: 64 Word.word) . - (undefined_vector (( 16 :: int)::ii) w__103 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__104 :: ( 64 bits) list) . + (undefined_vector (( 16 :: int)::ii) w__103 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__104 :: ( 64 Word.word) list) . (write_reg DBGWVR_EL1_ref w__104 \<then> (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__105 :: 32 Word.word) . - (undefined_vector (( 16 :: int)::ii) w__105 :: ( ( 32 Word.word)list) M) \<bind> (\<lambda> (w__106 :: ( 32 bits) list) . + (undefined_vector (( 16 :: int)::ii) w__105 :: ( ( 32 Word.word)list) M) \<bind> (\<lambda> (w__106 :: ( 32 Word.word) list) . (write_reg DBGWCR_EL1_ref w__106 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__107 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__107 :: 32 Word.word) . (write_reg DBGPRCR_EL1_ref w__107 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__108 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__108 :: 32 Word.word) . (write_reg DBGPRCR_ref w__108 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__109 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__109 :: 32 Word.word) . (write_reg DBGOSLSR_ref w__109 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__110 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__110 :: 32 Word.word) . (write_reg DBGOSDLR_ref w__110 \<then> undefined_signal () ) \<bind> (\<lambda> (w__111 :: signal) . (write_reg DBGEN_ref w__111 \<then> (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__112 :: 64 Word.word) . - (undefined_vector (( 16 :: int)::ii) w__112 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__113 :: ( 64 bits) list) . + (undefined_vector (( 16 :: int)::ii) w__112 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__113 :: ( 64 Word.word) list) . (write_reg DBGBVR_EL1_ref w__113 \<then> (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__114 :: 32 Word.word) . - (undefined_vector (( 16 :: int)::ii) w__114 :: ( ( 32 Word.word)list) M) \<bind> (\<lambda> (w__115 :: ( 32 bits) list) . + (undefined_vector (( 16 :: int)::ii) w__114 :: ( ( 32 Word.word)list) M) \<bind> (\<lambda> (w__115 :: ( 32 Word.word) list) . (write_reg DBGBCR_EL1_ref w__115 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__116 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__116 :: 32 Word.word) . (write_reg CPTR_EL3_ref w__116 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__117 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__117 :: 32 Word.word) . (write_reg CPTR_EL2_ref w__117 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__118 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__118 :: 32 Word.word) . (write_reg CPACR_EL1_ref w__118 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__119 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__119 :: 32 Word.word) . (write_reg CONTEXTIDR_EL2_ref w__119 \<then> - (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__120 :: 32 bits) . + (undefined_bitvector (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__120 :: 32 Word.word) . (write_reg CONTEXTIDR_EL1_ref w__120 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__121 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__121 :: 64 Word.word) . (write_reg APIBKeyLo_EL1_ref w__121 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__122 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__122 :: 64 Word.word) . (write_reg APIBKeyHi_EL1_ref w__122 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__123 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__123 :: 64 Word.word) . (write_reg APIAKeyLo_EL1_ref w__123 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__124 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__124 :: 64 Word.word) . (write_reg APIAKeyHi_EL1_ref w__124 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__125 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__125 :: 64 Word.word) . (write_reg APGAKeyLo_EL1_ref w__125 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__126 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__126 :: 64 Word.word) . (write_reg APGAKeyHi_EL1_ref w__126 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__127 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__127 :: 64 Word.word) . (write_reg APDBKeyLo_EL1_ref w__127 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__128 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__128 :: 64 Word.word) . (write_reg APDBKeyHi_EL1_ref w__128 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__129 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__129 :: 64 Word.word) . (write_reg APDAKeyLo_EL1_ref w__129 \<then> - (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__130 :: 64 bits) . + (undefined_bitvector (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__130 :: 64 Word.word) . write_reg APDAKeyHi_EL1_ref w__130))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))" diff --git a/snapshots/isabelle/aarch64/Aarch64_extras.thy b/snapshots/isabelle/aarch64/Aarch64_extras.thy index 7555341b..05647e36 100644 --- a/snapshots/isabelle/aarch64/Aarch64_extras.thy +++ b/snapshots/isabelle/aarch64/Aarch64_extras.thy @@ -1,15 +1,15 @@ -chapter \<open>Generated by Lem from ../aarch64_extras.lem.\<close> +chapter \<open>Generated by Lem from \<open>/auto/homes/tb592/REMS/sail/aarch64/mono/aarch64_extras.lem\<close>.\<close> theory "Aarch64_extras" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_operators_mwords" - "Sail2_prompt_monad" - "Sail2_prompt" +imports + Main + "LEM.Lem_pervasives_extra" + "Sail.Sail2_instr_kinds" + "Sail.Sail2_values" + "Sail.Sail2_operators_mwords" + "Sail.Sail2_prompt_monad" + "Sail.Sail2_prompt" begin diff --git a/snapshots/isabelle/aarch64/Aarch64_types.thy b/snapshots/isabelle/aarch64/Aarch64_types.thy index f233a3c0..d118f823 100644 --- a/snapshots/isabelle/aarch64/Aarch64_types.thy +++ b/snapshots/isabelle/aarch64/Aarch64_types.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from aarch64_types.lem.\<close> +chapter \<open>Generated by Lem from \<open>aarch64_types.lem\<close>.\<close> theory "Aarch64_types" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_operators_mwords" - "Sail2_prompt_monad" - "Sail2_prompt" - "Sail2_string" +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" begin diff --git a/snapshots/isabelle/cheri/Cheri.thy b/snapshots/isabelle/cheri/Cheri.thy index af6852b6..e4f28651 100644 --- a/snapshots/isabelle/cheri/Cheri.thy +++ b/snapshots/isabelle/cheri/Cheri.thy @@ -1,18 +1,18 @@ -chapter \<open>Generated by Lem from cheri.lem.\<close> +chapter \<open>Generated by Lem from \<open>cheri.lem\<close>.\<close> theory "Cheri" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_string" - "Sail2_operators_mwords" - "Sail2_prompt_monad" - "Sail2_prompt" - "Cheri_types" - "Mips_extras" +imports + Main + "LEM.Lem_pervasives_extra" + "Sail.Sail2_instr_kinds" + "Sail.Sail2_values" + "Sail.Sail2_string" + "Sail.Sail2_operators_mwords" + "Sail.Sail2_prompt_monad" + "Sail.Sail2_prompt" + "Cheri_types" + "Mips_extras" begin @@ -31,6 +31,12 @@ definition cap_size :: " int " where " cap_size = ( (( 32 :: int)::ii))" +(*val eq_unit : unit -> unit -> bool*) + +definition eq_unit :: " unit \<Rightarrow> unit \<Rightarrow> bool " where + " eq_unit g__20 g__21 = ( True )" + + @@ -48,7 +54,7 @@ definition neq_bool :: " bool \<Rightarrow> bool \<Rightarrow> bool " where definition undefined_option :: " 'a \<Rightarrow>((register_value),('a option),(exception))monad " where " undefined_option typ_a = ( undefined_unit () \<bind> (\<lambda> (u_0 :: unit) . - (let (u_1 :: 'a) = typ_a in + (let u_1 = typ_a in internal_pick [Some u_1,None])))" @@ -94,7 +100,8 @@ definition MIPS_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('p8_tim (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] - :: 64 Word.word) addr data )" + :: 64 Word.word) addr data \<then> + return () )" (*val __MIPS_read : forall 'p8_times_n_ . Size 'p8_times_n_ => mword ty64 -> integer -> M (mword 'p8_times_n_)*) @@ -1135,20 +1142,20 @@ definition TLBIndexMax :: "(6)Word.word " where (*val MAX : integer -> integer*) -definition MAX :: " int \<Rightarrow> int " where - " MAX n = ( ((pow2 n)) - (( 1 :: int)::ii))" +definition MAX0 :: " int \<Rightarrow> int " where + " MAX0 n = ( ((pow2 n)) - (( 1 :: int)::ii))" definition MAX_U64 :: " int " where - " MAX_U64 = ( MAX (( 64 :: int)::ii))" + " MAX_U64 = ( MAX0 (( 64 :: int)::ii))" definition MAX_VA :: " int " where - " MAX_VA = ( MAX (( 40 :: int)::ii))" + " MAX_VA = ( MAX0 (( 40 :: int)::ii))" definition MAX_PA :: " int " where - " MAX_PA = ( MAX (( 36 :: int)::ii))" + " MAX_PA = ( MAX0 (( 36 :: int)::ii))" (*val undefined_TLBEntry : unit -> M TLBEntry*) @@ -2203,7 +2210,7 @@ definition wGPR :: "(5)Word.word \<Rightarrow>(64)Word.word \<Rightarrow>((regi " wGPR idx v = ( (let i = (Word.uint idx) in if (((i \<noteq> (( 0 :: int)::ii)))) then - read_reg GPR_ref \<bind> (\<lambda> (w__0 :: ( 64 Word.word) list) . + read_reg GPR_ref \<bind> (\<lambda> (w__0 :: ( 64 bits) list) . write_reg GPR_ref ((update_list_dec w__0 i v :: ( 64 Word.word) list))) else return () ))" @@ -2317,7 +2324,7 @@ definition SignalExceptionMIPS :: " Exception \<Rightarrow>(64)Word.word \<Righ ((if ((\<not> ((bits_to_bool ((get_StatusReg_EXL w__0 :: 1 Word.word)))))) then (read_reg inBranchDelay_ref :: ( 1 Word.word) M) \<bind> (\<lambda> (w__1 :: 1 bits) . if ((bit_to_bool ((access_vec_dec w__1 (( 0 :: int)::ii))))) then - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . write_reg CP0EPC_ref ((sub_vec_int w__2 (( 4 :: int)::ii) :: 64 Word.word)) \<then> set_CauseReg_BD CP0Cause_ref (vec_of_bits [B1] :: 1 Word.word)) else @@ -2503,7 +2510,7 @@ definition SignalException :: " Exception \<Rightarrow>((register_value),'o,(ex read_reg CP0Status_ref \<bind> (\<lambda> (w__0 :: StatusReg) . ((if ((\<not> ((bits_to_bool ((get_StatusReg_EXL w__0 :: 1 Word.word)))))) then (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> pc . - (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__1 :: 257 Word.word) . + (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__1 :: CapReg) . (let pcc = (capRegToCapStruct w__1) in (let (success, epcc) = (setCapOffset pcc pc) in if success then write_reg EPCC_ref ((capStructToCapReg epcc :: 257 Word.word)) @@ -2522,7 +2529,7 @@ definition SignalException :: " Exception \<Rightarrow>((register_value),'o,(ex (write_reg nextPCC_ref w__2 \<then> (read_reg KCC_ref :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__3 :: CapReg) . (write_reg delayedPCC_ref w__3 \<then> - (read_reg KCC_ref :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__4 :: 257 Word.word) . + (read_reg KCC_ref :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__4 :: CapReg) . (let base = (getCapBase ((capRegToCapStruct w__4))) in SignalExceptionMIPS ex ((to_bits ((make_the_value (( 64 :: int)::ii) :: 64 itself)) base :: 64 Word.word))))))))" @@ -2593,7 +2600,7 @@ definition undefined_AccessLevel :: " unit \<Rightarrow>((register_value),(Acce " undefined_AccessLevel _ = ( internal_pick [User,Supervisor,Kernel])" -(*val int_of_AccessLevel : AccessLevel -> ii*) +(*val int_of_AccessLevel : AccessLevel -> integer*) fun int_of_AccessLevel :: " AccessLevel \<Rightarrow> int " where " int_of_AccessLevel User = ( (( 0 :: int)::ii))" @@ -2626,8 +2633,8 @@ definition getAccessLevel :: " unit \<Rightarrow>((register_value),(AccessLevel if w__2 then return Kernel else read_reg CP0Status_ref \<bind> (\<lambda> (w__3 :: StatusReg) . - (let p__31 = ((get_StatusReg_KSU w__3 :: 2 Word.word)) in - (let b__0 = p__31 in + (let p__19 = ((get_StatusReg_KSU w__3 :: 2 Word.word)) in + (let b__0 = p__19 in return (if (((b__0 = (vec_of_bits [B0,B0] :: 2 Word.word)))) then Kernel else if (((b__0 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then Supervisor else if (((b__0 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then User @@ -2653,16 +2660,16 @@ definition checkCP0Access :: " unit \<Rightarrow>((register_value),(unit),(exce definition incrementCP0Count :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " incrementCP0Count _ = ( (read_reg TLBRandom_ref :: ( 6 Word.word) M) \<bind> (\<lambda> (w__0 :: TLBIndexT) . - (read_reg TLBWired_ref :: ( 6 Word.word) M) \<bind> (\<lambda> (w__1 :: 6 Word.word) . + (read_reg TLBWired_ref :: ( 6 Word.word) M) \<bind> (\<lambda> (w__1 :: TLBIndexT) . (if (((w__0 = w__1))) then return TLBIndexMax else - (read_reg TLBRandom_ref :: ( 6 Word.word) M) \<bind> (\<lambda> (w__2 :: 6 Word.word) . + (read_reg TLBRandom_ref :: ( 6 Word.word) M) \<bind> (\<lambda> (w__2 :: TLBIndexT) . return ((sub_vec_int w__2 (( 1 :: int)::ii) :: 6 Word.word)))) \<bind> (\<lambda> (w__3 :: 6 Word.word) . (write_reg TLBRandom_ref w__3 \<then> - (read_reg CP0Count_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__4 :: 32 Word.word) . + (read_reg CP0Count_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__4 :: 32 bits) . (write_reg CP0Count_ref ((add_vec_int w__4 (( 1 :: int)::ii) :: 32 Word.word)) \<then> (read_reg CP0Count_ref :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__5 :: 32 bits) . - (read_reg CP0Compare_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 Word.word) . + (read_reg CP0Compare_ref :: ( 32 Word.word) M) \<bind> (\<lambda> (w__6 :: 32 bits) . ((if (((w__5 = w__6))) then read_reg CP0Cause_ref \<bind> (\<lambda> (w__7 :: CauseReg) . set_CauseReg_IP CP0Cause_ref @@ -2907,7 +2914,7 @@ definition extsv :: " int \<Rightarrow>('n::len)Word.word \<Rightarrow>('m::len definition slice_mask :: " int \<Rightarrow> int \<Rightarrow> int \<Rightarrow>('n::len)Word.word " where " slice_mask (n__tv :: int) i l = ( - (let (one :: 'n bits) = ((extzv n__tv (vec_of_bits [B1] :: 1 Word.word) :: ( 'n::len)Word.word)) in + (let one = ((extzv n__tv (vec_of_bits [B1] :: 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)))" @@ -2924,8 +2931,7 @@ definition is_zero_subrange :: "('n::len)Word.word \<Rightarrow> int \<Rightarr definition is_ones_subrange :: "('n::len)Word.word \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool " where " is_ones_subrange xs i j = ( - (let (m :: 'n bits) = - ((slice_mask ((int (size xs))) j ((((j - i)) + (( 1 :: int)::ii))) :: ( 'n::len)Word.word)) in + (let m = ((slice_mask ((int (size xs))) j ((((j - i)) + (( 1 :: int)::ii))) :: ( 'n::len)Word.word)) in (((and_vec xs m :: ( 'n::len)Word.word)) = m)))" @@ -3062,7 +3068,7 @@ definition unsigned_subrange :: "('n::len)Word.word \<Rightarrow> int \<Rightar definition zext_ones :: " int \<Rightarrow> int \<Rightarrow>('n::len)Word.word " where " zext_ones (n__tv :: int) m = ( - (let (v :: 'n bits) = ((extsv n__tv (vec_of_bits [B1] :: 1 Word.word) :: ( 'n::len)Word.word)) in + (let v = ((extsv n__tv (vec_of_bits [B1] :: 1 Word.word) :: ( 'n::len)Word.word)) in (shiftr v ((((int (size v))) - m)) :: ( 'n::len)Word.word)))" @@ -3213,8 +3219,8 @@ definition TLBTranslateC :: "(64)Word.word \<Rightarrow> MemAccessType \<Righta ((subrange_vec_dec vAddr (( 28 :: int)::ii) (( 0 :: int)::ii) :: 29 Word.word)) :: 32 Word.word)) :: 64 Word.word))) - else (case (True, b__1) of (g__29, g__30) => (Kernel, None) ) - | (g__29, g__30) => (Kernel, None) + else (case (True, b__1) of (g__17, g__18) => (Kernel, None) ) + | (g__17, g__18) => (Kernel, None) ) else if (((b__0 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then (Kernel, @@ -3315,13 +3321,13 @@ definition undefined_CapStruct :: " unit \<Rightarrow>((register_value),(CapStr " undefined_CapStruct _ = ( undefined_bool () \<bind> (\<lambda> (w__0 :: bool) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 8 :: int)::ii) :: ( 8 Word.word) M) \<bind> (\<lambda> (w__1 :: 8 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 24 :: int)::ii) :: ( 24 Word.word) M) \<bind> (\<lambda> (w__2 :: 24 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 24 :: int)::ii) :: ( 24 Word.word) M) \<bind> (\<lambda> (w__2 :: 24 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (w__3 :: 16 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 16 :: int)::ii) :: ( 16 Word.word) M) \<bind> (\<lambda> (w__3 :: 16 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__4 :: 4 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 4 :: int)::ii) :: ( 4 Word.word) M) \<bind> (\<lambda> (w__4 :: 4 Word.word) . undefined_bool () \<bind> (\<lambda> (w__5 :: bool) . undefined_bool () \<bind> (\<lambda> (w__6 :: bool) . undefined_bool () \<bind> (\<lambda> (w__7 :: bool) . @@ -3335,11 +3341,11 @@ definition undefined_CapStruct :: " unit \<Rightarrow>((register_value),(CapStr undefined_bool () \<bind> (\<lambda> (w__15 :: bool) . undefined_bool () \<bind> (\<lambda> (w__16 :: bool) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: 64 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__18 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__18 :: 64 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__19 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__19 :: 64 Word.word) . return ((| CapStruct_tag = w__0, CapStruct_padding = w__1, CapStruct_otype = w__2, @@ -3557,7 +3563,7 @@ definition CapRegs :: "(((regstate),(register_value),((257)Word.word))register_ definition max_otype :: " int " where - " max_otype = ( MAX (( 24 :: int)::ii))" + " max_otype = ( MAX0 (( 24 :: int)::ii))" definition have_cp2 :: " bool " where @@ -3854,7 +3860,7 @@ definition raise_c2_exception_noreg :: " CapEx \<Rightarrow>((register_value),' definition pcc_access_system_regs :: " unit \<Rightarrow>((register_value),(bool),(exception))monad " where " pcc_access_system_regs _ = ( - (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: 257 Word.word) . + (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: CapReg) . (let pcc = (capRegToCapStruct w__0) in return(CapStruct_access_system_regs pcc))))" @@ -3868,7 +3874,7 @@ definition register_inaccessible :: "(5)Word.word \<Rightarrow>((register_value " register_inaccessible r = ( or_boolM (and_boolM (return (((r = IDCNO)))) - ((read_reg inCCallDelay_ref :: ( 1 Word.word) M) \<bind> (\<lambda> (w__0 :: 1 Word.word) . + ((read_reg inCCallDelay_ref :: ( 1 Word.word) M) \<bind> (\<lambda> (w__0 :: 1 bits) . return ((bits_to_bool w__0))))) (and_boolM (return ((((((r = KR1CNO))) \<or> ((((((r = KR2CNO))) \<or> ((((((r = KDCNO))) \<or> ((((((r = KCCNO))) \<or> (((r = EPCCNO)))))))))))))))) @@ -3993,7 +3999,7 @@ definition checkDDCPerms :: " CapStruct \<Rightarrow> MemAccessType \<Rightarro definition addrWrapper :: "(64)Word.word \<Rightarrow> MemAccessType \<Rightarrow> WordType \<Rightarrow>((register_value),((64)Word.word),(exception))monad " where " addrWrapper addr accessType width = ( - (read_reg DDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: 257 Word.word) . + (read_reg DDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: CapReg) . (let ddc = (capRegToCapStruct w__0) in checkDDCPerms ddc accessType \<then> ((let cursor = (getCapCursor ddc) in @@ -4014,7 +4020,7 @@ definition addrWrapper :: "(64)Word.word \<Rightarrow> MemAccessType \<Rightarr definition addrWrapperUnaligned :: "(64)Word.word \<Rightarrow> MemAccessType \<Rightarrow> WordTypeUnaligned \<Rightarrow>((register_value),((64)Word.word),(exception))monad " where " addrWrapperUnaligned addr accessType width = ( - (read_reg DDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: 257 Word.word) . + (read_reg DDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: CapReg) . (let ddc = (capRegToCapStruct w__0) in checkDDCPerms ddc accessType \<then> ((let cursor = (getCapCursor ddc) in @@ -4044,7 +4050,7 @@ definition addrWrapperUnaligned :: "(64)Word.word \<Rightarrow> MemAccessType \ definition TranslatePC :: "(64)Word.word \<Rightarrow>((register_value),((64)Word.word),(exception))monad " where " TranslatePC vAddr = ( (incrementCP0Count () \<then> - (read_reg PCC_ref :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__0 :: 257 Word.word) . + (read_reg PCC_ref :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__0 :: CapReg) . (let pcc = (capRegToCapStruct w__0) in (let base = (getCapBase pcc) in (let top1 = (getCapTop pcc) in @@ -4105,7 +4111,7 @@ definition cp2_next_pc :: " unit \<Rightarrow>((register_value),(unit),(excepti " cp2_next_pc _ = ( (read_reg nextPCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: CapReg) . (write_reg PCC_ref w__0 \<then> - (read_reg inBranchDelay_ref :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__1 :: 1 Word.word) . + (read_reg inBranchDelay_ref :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__1 :: 1 bits) . if ((bits_to_bool w__1)) then (read_reg delayedPCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__2 :: CapReg) . write_reg nextPCC_ref w__2) @@ -4117,81 +4123,81 @@ definition cp2_next_pc :: " unit \<Rightarrow>((register_value),(unit),(excepti definition capToString :: " CapStruct \<Rightarrow>((register_value),(string),(exception))monad " where " capToString cap = ( skip () \<then> - return (((op@) ('' t:'') - (((op@) (if(CapStruct_tag cap) then (''1'') else (''0'')) - (((op@) ('' s:'') - (((op@) (if(CapStruct_sealed cap) then (''1'') else (''0'')) - (((op@) ('' perms:'') - (((op@) + return (((@) ('' t:'') + (((@) (if(CapStruct_tag cap) then (''1'') else (''0'')) + (((@) ('' s:'') + (((@) (if(CapStruct_sealed cap) then (''1'') else (''0'')) + (((@) ('' perms:'') + (((@) ((string_of_bits ((concat_vec (vec_of_bits [B0] :: 1 Word.word) ((getCapPerms cap :: 31 Word.word)) :: 32 Word.word)))) - (((op@) ('' type:'') - (((op@) ((string_of_bits(CapStruct_otype cap))) - (((op@) ('' offset:'') - (((op@) + (((@) ('' type:'') + (((@) ((string_of_bits(CapStruct_otype cap))) + (((@) ('' offset:'') + (((@) ((string_of_bits ((to_bits ((make_the_value (( 64 :: int)::ii) :: 64 itself)) ((getCapOffset cap)) :: 64 Word.word)))) - (((op@) ('' base:'') - (((op@) + (((@) ('' base:'') + (((@) ((string_of_bits ((to_bits ((make_the_value (( 64 :: int)::ii) :: 64 itself)) ((getCapBase cap)) :: 64 Word.word)))) - (((op@) ('' length:'') + (((@) ('' length:'') ((string_of_bits ((to_bits ((make_the_value (( 64 :: int)::ii) :: 64 itself)) ((min ((getCapLength cap)) - ((MAX (( 64 :: int)::ii))))) + ((MAX0 (( 64 :: int)::ii))))) :: 64 Word.word)))))))))))))))))))))))))))))))" definition dump_cp2_state :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " dump_cp2_state _ = ( - (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: 257 Word.word) . + (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__0 :: CapReg) . capToString ((capRegToCapStruct w__0)) \<bind> (\<lambda> (w__1 :: string) . - (let (_ :: unit) = (print_endline (((op@) (''DEBUG CAP PCC'') w__1))) in + (let (_ :: unit) = (print_endline (((@) (''DEBUG CAP PCC'') w__1))) in ((foreachM (index_list (( 0 :: int)::ii) (( 31 :: int)::ii) (( 1 :: int)::ii)) () (\<lambda> i unit_var . readCapReg ((to_bits ((make_the_value (( 5 :: int)::ii) :: 5 itself)) i :: 5 Word.word)) \<bind> (\<lambda> (w__2 :: CapStruct) . capToString w__2 \<bind> (\<lambda> (w__3 :: string) . return ((let _ = - (print_endline (((op@) (''DEBUG CAP REG '') (((op@) ((string_of_int + (print_endline (((@) (''DEBUG CAP REG '') (((@) ((string_of_int instance_Show_Show_Num_integer_dict i)) w__3))))) in () )))))) \<then> - (read_reg DDC_ref :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__4 :: 257 Word.word) . + (read_reg DDC_ref :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__4 :: CapReg) . capToString ((capRegToCapStruct w__4)) \<bind> (\<lambda> (w__5 :: string) . - (let (_ :: unit) = (print_endline (((op@) (''DEBUG CAP HWREG 00'') w__5))) in - (read_reg CTLSU_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__6 :: 257 Word.word) . + (let (_ :: unit) = (print_endline (((@) (''DEBUG CAP HWREG 00'') w__5))) in + (read_reg CTLSU_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__6 :: CapReg) . capToString ((capRegToCapStruct w__6)) \<bind> (\<lambda> (w__7 :: string) . - (let (_ :: unit) = (print_endline (((op@) (''DEBUG CAP HWREG 01'') w__7))) in - (read_reg CTLSP_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__8 :: 257 Word.word) . + (let (_ :: unit) = (print_endline (((@) (''DEBUG CAP HWREG 01'') w__7))) in + (read_reg CTLSP_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__8 :: CapReg) . capToString ((capRegToCapStruct w__8)) \<bind> (\<lambda> (w__9 :: string) . - (let (_ :: unit) = (print_endline (((op@) (''DEBUG CAP HWREG 08'') w__9))) in - (read_reg KR1C_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__10 :: 257 Word.word) . + (let (_ :: unit) = (print_endline (((@) (''DEBUG CAP HWREG 08'') w__9))) in + (read_reg KR1C_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__10 :: CapReg) . capToString ((capRegToCapStruct w__10)) \<bind> (\<lambda> (w__11 :: string) . - (let (_ :: unit) = (print_endline (((op@) (''DEBUG CAP HWREG 22'') w__11))) in - (read_reg KR2C_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__12 :: 257 Word.word) . + (let (_ :: unit) = (print_endline (((@) (''DEBUG CAP HWREG 22'') w__11))) in + (read_reg KR2C_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__12 :: CapReg) . capToString ((capRegToCapStruct w__12)) \<bind> (\<lambda> (w__13 :: string) . - (let (_ :: unit) = (print_endline (((op@) (''DEBUG CAP HWREG 23'') w__13))) in - (read_reg KCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__14 :: 257 Word.word) . + (let (_ :: unit) = (print_endline (((@) (''DEBUG CAP HWREG 23'') w__13))) in + (read_reg KCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__14 :: CapReg) . capToString ((capRegToCapStruct w__14)) \<bind> (\<lambda> (w__15 :: string) . - (let (_ :: unit) = (print_endline (((op@) (''DEBUG CAP HWREG 29'') w__15))) in - (read_reg KDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__16 :: 257 Word.word) . + (let (_ :: unit) = (print_endline (((@) (''DEBUG CAP HWREG 29'') w__15))) in + (read_reg KDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__16 :: CapReg) . capToString ((capRegToCapStruct w__16)) \<bind> (\<lambda> (w__17 :: string) . - (let (_ :: unit) = (print_endline (((op@) (''DEBUG CAP HWREG 30'') w__17))) in - (read_reg EPCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__18 :: 257 Word.word) . + (let (_ :: unit) = (print_endline (((@) (''DEBUG CAP HWREG 30'') w__17))) in + (read_reg EPCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__18 :: CapReg) . capToString ((capRegToCapStruct w__18)) \<bind> (\<lambda> (w__19 :: string) . - return ((print_endline (((op@) (''DEBUG CAP HWREG 31'') w__19)))))))))))))))))))))))))))))))" + return ((print_endline (((@) (''DEBUG CAP HWREG 31'') w__19)))))))))))))))))))))))))))))))" (*val extendLoad : forall 'sz . Size 'sz => mword 'sz -> bool -> mword ty64*) @@ -4273,98 +4279,98 @@ definition TLBWriteEntry :: "(6)Word.word \<Rightarrow>((register_value),(unit) definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where " decode v__0 = ( if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (DADDIU (rs,rt,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B1,B1,B0,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (DADDU (rs,rt,rd))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B0] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (DADDI (rs,rt,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B1,B1,B0,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (DADD (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (ADD (rs,rt,rd))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B0] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (ADDI (rs,rt,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (ADDU (rs,rt,rd))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (ADDIU (rs,rt,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B1,B1,B1,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (DSUBU (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B1,B1,B1,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (DSUB (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B1,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (SUB (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B1,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (SUBU (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B1,B0,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (AND0 (rs,rt,rd))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B0] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (ANDI (rs,rt,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B1,B0,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (OR0 (rs,rt,rd))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (ORI (rs,rt,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B1,B1,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (NOR (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B1,B1,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (XOR0 (rs,rt,rd))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B1,B0] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (XORI (rs,rt,imm))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B1,B1,B0,B0,B0,B0,B0] :: 11 Word.word)))) then @@ -4372,108 +4378,108 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (LUI (rt,imm)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B0,B0] :: 6 Word.word))))))) then + (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (DSLL (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B1,B0,B0] :: 6 Word.word))))))) then + (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (DSLL32 (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1,B0,B1,B0,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (DSLLV (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B1,B1] :: 6 Word.word))))))) then + (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (DSRA (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B1,B1,B1] :: 6 Word.word))))))) then + (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (DSRA32 (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1,B0,B1,B1,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (DSRAV (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B1,B0] :: 6 Word.word))))))) then + (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (DSRL (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B1,B1,B0] :: 6 Word.word))))))) then + (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: 5 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (DSRL32 (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1,B0,B1,B1,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (DSRLV (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then + (let (sa :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (SLL (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (SLLV (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B1,B1] :: 6 Word.word))))))) then + (let (sa :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (SRA (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (SRAV (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B1,B0] :: 6 Word.word))))))) then + (let (sa :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sa :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (SRL (rt,rd,sa))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (SRLV (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B1,B0,B1,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (SLT (rs,rt,rd))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B1,B0] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (SLTI (rs,rt,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B1,B0,B1,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (SLTU (rs,rt,rd))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B1,B1] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (SLTIU (rs,rt,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B1,B0,B1,B1] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (MOVN (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B1,B0,B1,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (MOVZ (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 16 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0] :: 11 Word.word))))))) then @@ -4491,69 +4497,69 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (MTLO rs)) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 11 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (MUL (rs,rt,rd))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B0,B0,B0] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (MULT (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B0,B0,B1] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (MULTU (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B0,B0] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (DMULT (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B0,B1] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (DMULTU (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (MADD (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (MADDU (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (MSUB (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B1] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (MSUBU (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B0,B1,B0] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (DIV (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B0,B1,B1] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (DIVU (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B1,B0] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (DDIV (rs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B1,B1] :: 16 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (DDIVU (rs,rt)))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B1,B0] :: 6 Word.word)))) then (let (offset :: 26 bits) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 0 :: int)::ii) :: 26 Word.word)) in @@ -4569,23 +4575,23 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (JALR (rs,rd)))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B0,B0] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (BEQ (rs,rt,imm,False,False))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B0,B0] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (BEQ (rs,rt,imm,False,True))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B0,B1] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (BEQ (rs,rt,imm,True,False))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B0,B1] :: 6 Word.word)))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (BEQ (rs,rt,imm,True,True))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))) then @@ -4645,28 +4651,28 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where :: 32 Word.word)))) then Some (WAIT () ) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B0,B0] :: 6 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (TRAPREG (rs,rt,GE)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B0,B1] :: 6 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (TRAPREG (rs,rt,GEU)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B0] :: 6 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (TRAPREG (rs,rt,LT')))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B1] :: 6 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (TRAPREG (rs,rt,LTU)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B1,B0,B0] :: 6 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (TRAPREG (rs,rt,EQ')))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B1,B1,B0] :: 6 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (TRAPREG (rs,rt,NE)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word))))))) then (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in @@ -4693,137 +4699,137 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (TRAPIMM (rs,imm,LTU)))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B0,B0,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (B,True,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B1,B0,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (B,False,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B0,B0,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (H,True,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B1,B0,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (H,False,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B0,B1,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (W,True,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B1,B1,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (W,False,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B1,B1,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (D,False,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B0,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (W,True,True,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B1,B0,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Load (D,False,True,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B1,B0,B0,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Store (B,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B1,B0,B0,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Store (H,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B1,B0,B1,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Store (W,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B1,B1,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Store (D,False,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B0,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Store (W,True,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B1,B0,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (Store (D,True,base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B0,B1,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (LWL (base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B1,B1,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (LWR (base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B1,B0,B1,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (SWL (base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B1,B1,B1,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (SWR (base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B1,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (LDL (base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B1,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (LDR (base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B1,B1,B0,B0] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (SDL (base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B1,B1,B0,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (offset :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (SDR (base,rt,offset))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B1,B1,B1,B1] :: 6 Word.word)))) then - (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (op1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (imm :: imm16) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (base :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in Some (CACHE (base,op1,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 11 :: int)::ii) :: 21 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 21 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B1,B1] :: 6 Word.word))))))) then Some (SYNC () ) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))))))) then + (let (sel :: 3 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sel :: 3 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in Some (MFC0 (rt,rd,sel,False))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))))))) then + (let (sel :: 3 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sel :: 3 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in Some (MFC0 (rt,rd,sel,True))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 16 Word.word))))))) then @@ -4832,14 +4838,14 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where then Some (HCF () ) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))))))) then + (let (sel :: 3 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sel :: 3 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in Some (MTC0 (rt,rd,sel,False))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0,B0,B1,B0,B1] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))))))) then + (let (sel :: 3 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (sel :: 3 bits) = ((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) in Some (MTC0 (rt,rd,sel,True))))) else if (((v__0 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, B0,B0,B0,B0,B0,B0,B1,B0] @@ -4906,96 +4912,96 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CSetCause rt)) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then + (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CAndPerm (cd1,cb,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CToPtr (rd,cb,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,ct,CEQ))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,ct,CNE))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B1,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,ct,CLT))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B1,B1] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,ct,CLE))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B0,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,ct,CLTU))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B0,B1] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,ct,CLEU))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B1,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,ct,CEXEQ))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B1,B1] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,ct,CNEXEQ))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B0,B1] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then + (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CIncOffset (cd1,cb,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B0,B1] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1] :: 6 Word.word))))))) then + (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CSetOffset (cd1,cb,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B1] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then + (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CSetBounds (cd1,cb,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B1] :: 11 Word.word))))))) then (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CClearTag (cd1,cb)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B1,B1] :: 6 Word.word))))))) then + (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CFromPtr (cd1,cb,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B0,B1,B1] :: 11 Word.word)))) \<and> ((((((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))))))))) then - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CCheckPerm (cs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B0,B1,B1] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 11 Word.word))))))) then (let (cs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CCheckType (cs,cb)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B1,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CSeal (cd1,cs,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B1,B1] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CUnseal (cd1,cs,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B1,B1,B1] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 11 Word.word))))))) then (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in @@ -5021,8 +5027,8 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CJALR ((vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),cb,False))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B1,B1,B1,B1,B1,B1] :: 11 Word.word))))))) then - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CCheckPerm (cs,rt)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B1,B1,B1,B1,B1,B1] :: 11 Word.word))))))) then (let (cs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in @@ -5033,8 +5039,8 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CClearTag (cd1,cb)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B0,B1,B1,B1,B1,B1,B1] :: 11 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CMOVX (cd1,cs,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),False)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1,B1,B1,B1,B1] :: 11 Word.word))))))) then (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in @@ -5069,165 +5075,165 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CGetOffset (rd,cb)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B1,B1,B1,B1,B1,B1,B1] :: 11 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CGetPCCSetOffset (cd1,rs)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B1,B1,B1,B1,B1,B1,B1] :: 11 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (sel :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CReadHwr (cd1,sel)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B1] :: 11 Word.word))))))) then - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (sel :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CWriteHwr (cb,sel)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B1,B1,B1,B1,B1,B1,B1] :: 11 Word.word))))))) then - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (sel :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CGetAddr (cb,sel)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B1,B1] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CSeal (cd1,cs,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B0] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CUnseal (cd1,cs,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CAndPerm (cd1,cs,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B1,B1] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CSetOffset (cd1,cs,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B0] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CSetBounds (cd1,cs,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CSetBoundsExact (cd1,cs,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B1] :: 6 Word.word))))))) then + (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (rt :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CIncOffset (cd1,cb,rt))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1] :: 6 Word.word))))))) then + (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CBuildCap (cd1,cb,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B1,B0] :: 6 Word.word))))))) then + (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CCopyType (cd1,cb,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B1,B1] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CCSeal (cd1,cs,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CToPtr (rd,cb,ct))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B1] :: 6 Word.word))))))) then + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CFromPtr (cd1,cb,rs))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B1,B0] :: 6 Word.word))))))) then (let (rt :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CSub (rt,cb,cs))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B1,B1] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CMOVX (cd1,cs,rs,False))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B0] :: 6 Word.word))))))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CMOVX (cd1,cs,rs,True))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B0,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,cs,CEQ))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B0,B1] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,cs,CNE))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B1,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,cs,CLT))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B1,B1] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,cs,CLE))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,cs,CLTU))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,cs,CLEU))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B1,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,cs,CEXEQ))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B0,B0,B1] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CPtrCmp (rd,cb,cs,CNEXEQ))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (ct :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CTestSubset (rd,cb,ct))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B0,B0,B1] :: 11 Word.word)))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (imm :: 16 bits) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CBX (cd1,imm,True)))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B0,B1,B0] :: 11 Word.word)))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (imm :: 16 bits) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CBX (cd1,imm,False)))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B0,B1] :: 11 Word.word)))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (imm :: 16 bits) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CBZ (cd1,imm,False)))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B1,B0] :: 11 Word.word)))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (imm :: 16 bits) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CBZ (cd1,imm,True)))) else if (((v__0 = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1, B1,B1,B1,B1,B1,B1,B1,B1] :: 32 Word.word)))) then Some (CReturn () ) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B1,B0,B1] :: 11 Word.word)))) then + (let (selector :: 11 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (selector :: 11 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) in Some (CCall (cs,cb,selector))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B1,B1,B1,B1,B0,B0,B0,B0,B0] :: 16 Word.word)))) then (let (imm :: 16 bits) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in @@ -5242,56 +5248,56 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where (let (imm :: 16 bits) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (ClearRegs (CHi,imm))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B1,B1] :: 11 Word.word)))) then + (let (imm :: 11 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (imm :: 11 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) in Some (CIncOffsetImmediate (cd1,cb,imm))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B1,B0,B0] :: 11 Word.word)))) then + (let (imm :: 11 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (imm :: 11 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) in Some (CSetBoundsImmediate (cd1,cb,imm))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) then - (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CLoad (rd,cb,rt,offset,False,B,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B0] :: 3 Word.word))))))) then - (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CLoad (rd,cb,rt,offset,True,B,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) then - (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CLoad (rd,cb,rt,offset,False,H,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) then - (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CLoad (rd,cb,rt,offset,True,H,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B0] :: 3 Word.word))))))) then - (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CLoad (rd,cb,rt,offset,False,W,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B0] :: 3 Word.word))))))) then - (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CLoad (rd,cb,rt,offset,True,W,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word))))))) then - (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CLoad (rd,cb,rt,offset,False,D,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0] :: 11 Word.word))))))) then (let (rd :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in @@ -5336,75 +5342,75 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where B0,B0] :: 8 Word.word),False,D,True)))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CStore (rs,cb,rt,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),offset,B,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CStore (rs,cb,rt,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),offset,H,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B0] :: 3 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CStore (rs,cb,rt,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),offset,W,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B1,B0] :: 6 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word))))))) then - (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in + (let (rs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (offset :: 8 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 3 :: int)::ii) :: 8 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CStore (rs,cb,rt,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),offset,D,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))) then (let (rs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CStore (rs,cb,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),rd,(vec_of_bits [B0,B0,B0,B0,B0, B0,B0,B0] :: 8 Word.word),B,True))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1] :: 6 Word.word))))))) then (let (rs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CStore (rs,cb,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),rd,(vec_of_bits [B0,B0,B0,B0,B0, B0,B0,B0] :: 8 Word.word),H,True))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B1,B0] :: 6 Word.word))))))) then (let (rs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CStore (rs,cb,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),rd,(vec_of_bits [B0,B0,B0,B0,B0, B0,B0,B0] :: 8 Word.word),W,True))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B1,B1] :: 6 Word.word))))))) then (let (rs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (rd :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in Some (CStore (rs,cb,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),rd,(vec_of_bits [B0,B0,B0,B0,B0, B0,B0,B0] :: 8 Word.word),D,True))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B1,B1,B1,B0] :: 6 Word.word)))) then - (let (cs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (offset :: 11 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) in + (let (cs :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CSC (cs,cb,rt,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),offset,False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 5 :: int)::ii) (( 0 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B1,B1] :: 6 Word.word))))))) then + (let (rd :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in (let (cs :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in - (let (rd :: regno) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 6 :: int)::ii) :: 5 Word.word)) in Some (CSC (cs,cb,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),rd,(vec_of_bits [B0,B0,B0,B0,B0,B0, B0,B0,B0,B0,B0] :: 11 Word.word),True))))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B1,B0,B1,B1,B0] :: 6 Word.word)))) then - (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in - (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in (let (rt :: regno) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 11 :: int)::ii) :: 5 Word.word)) in (let (offset :: 11 bits) = ((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) in + (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in + (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in Some (CLC (cd1,cb,rt,(mips_sign_extend (( 16 :: int)::ii) offset :: 16 Word.word),False)))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B1,B0,B0,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 10 :: int)::ii) (( 0 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B1] :: 11 Word.word))))))) then (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in @@ -5414,9 +5420,9 @@ definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where B0,B0] :: 16 Word.word),True)))) else if (((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1] :: 6 Word.word)))) then + (let (offset :: 16 bits) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in (let (cd1 :: regno) = ((subrange_vec_dec v__0 (( 25 :: int)::ii) (( 21 :: int)::ii) :: 5 Word.word)) in (let (cb :: regno) = ((subrange_vec_dec v__0 (( 20 :: int)::ii) (( 16 :: int)::ii) :: 5 Word.word)) in - (let (offset :: 16 bits) = ((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) in Some (CLC (cd1,cb,(vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word),offset,False))))) else if ((((((((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 21 :: int)::ii) :: 11 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 11 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 15 :: int)::ii) (( 0 :: int)::ii) :: 16 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B0] :: 16 Word.word))))))) then @@ -5445,7 +5451,7 @@ definition execute_XOR :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> (*val execute_WAIT : unit -> M unit*) definition execute_WAIT :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_WAIT g__19 = ( + " execute_WAIT _ = ( (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . write_reg nextPC_ref w__0))" @@ -5474,23 +5480,23 @@ definition execute_TRAPIMM :: "(5)Word.word \<Rightarrow>(16)Word.word \<Righta (*val execute_TLBWR : unit -> M unit*) definition execute_TLBWR :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_TLBWR g__23 = ( + " execute_TLBWR _ = ( (checkCP0Access () \<then> - (read_reg TLBRandom_ref :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__0 :: 6 Word.word) . TLBWriteEntry w__0))" + (read_reg TLBRandom_ref :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__0 :: TLBIndexT) . TLBWriteEntry w__0))" (*val execute_TLBWI : unit -> M unit*) definition execute_TLBWI :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_TLBWI g__22 = ( + " execute_TLBWI _ = ( (checkCP0Access () \<then> - (read_reg TLBIndex_ref :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__0 :: 6 Word.word) . TLBWriteEntry w__0))" + (read_reg TLBIndex_ref :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__0 :: TLBIndexT) . TLBWriteEntry w__0))" (*val execute_TLBR : unit -> M unit*) definition execute_TLBR :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_TLBR g__24 = ( + " execute_TLBR _ = ( (checkCP0Access () \<then> (read_reg TLBIndex_ref :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__0 :: TLBIndexT) . (let i = (Word.uint w__0) in @@ -5518,7 +5524,7 @@ definition execute_TLBR :: " unit \<Rightarrow>((register_value),(unit),(except (*val execute_TLBP : unit -> M unit*) definition execute_TLBP :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_TLBP g__25 = ( + " execute_TLBP _ = ( (checkCP0Access () \<then> read_reg TLBEntryHi_ref) \<bind> (\<lambda> (w__0 :: TLBEntryHiReg) . (tlbSearch ((get_TLBEntryHiReg_bits w__0 :: 64 Word.word)) :: ( ( 6 Word.word)option) M) \<bind> (\<lambda> result . @@ -5567,13 +5573,13 @@ definition execute_Store :: " WordType \<Rightarrow> bool \<Rightarrow>(5)Word. (*val execute_SYSCALL : unit -> M unit*) definition execute_SYSCALL :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_SYSCALL g__17 = ( SignalException Sys )" + " execute_SYSCALL _ = ( SignalException Sys )" (*val execute_SYNC : unit -> M unit*) definition execute_SYNC :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_SYNC g__20 = ( MEM_sync () )" + " execute_SYNC _ = ( MEM_sync () )" (*val execute_SWR : mword ty5 -> mword ty5 -> mword ty16 -> M unit*) @@ -5866,7 +5872,7 @@ definition execute_SDL :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> (*val execute_RI : unit -> M unit*) definition execute_RI :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_RI g__28 = ( SignalException ResI )" + " execute_RI _ = ( SignalException ResI )" (*val execute_RDHWR : mword ty5 -> mword ty5 -> M unit*) @@ -5996,14 +6002,14 @@ definition execute_MUL :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> definition execute_MTLO :: "(5)Word.word \<Rightarrow>((register_value),(unit),(exception))monad " where " execute_MTLO rs = ( - (rGPR rs :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . write_reg LO_ref w__0))" + (rGPR rs :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . write_reg LO_ref w__0))" (*val execute_MTHI : mword ty5 -> M unit*) definition execute_MTHI :: "(5)Word.word \<Rightarrow>((register_value),(unit),(exception))monad " where " execute_MTHI rs = ( - (rGPR rs :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . write_reg HI_ref w__0))" + (rGPR rs :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . write_reg HI_ref w__0))" (*val execute_MTC0 : mword ty5 -> mword ty5 -> mword ty3 -> bool -> M unit*) @@ -6191,15 +6197,13 @@ definition execute_MOVN :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow (*val execute_MFLO : mword ty5 -> M unit*) definition execute_MFLO :: "(5)Word.word \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_MFLO rd = ( - (read_reg LO_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . wGPR rd w__0))" + " execute_MFLO rd = ( (read_reg LO_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . wGPR rd w__0))" (*val execute_MFHI : mword ty5 -> M unit*) definition execute_MFHI :: "(5)Word.word \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_MFHI rd = ( - (read_reg HI_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . wGPR rd w__0))" + " execute_MFHI rd = ( (read_reg HI_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . wGPR rd w__0))" (*val execute_MFC0 : mword ty5 -> mword ty5 -> mword ty3 -> bool -> M unit*) @@ -6675,7 +6679,7 @@ definition execute_JALR :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow " execute_JALR rs rd = ( (rGPR rs :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . (execute_branch w__0 \<then> - (read_reg PC_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 bits) . wGPR rd ((add_vec_int w__1 (( 8 :: int)::ii) :: 64 Word.word)))))" @@ -6689,7 +6693,7 @@ definition execute_JAL :: "(26)Word.word \<Rightarrow>((register_value),(unit), ((subrange_vec_dec ((add_vec_int w__0 (( 4 :: int)::ii) :: 64 Word.word)) (( 63 :: int)::ii) (( 28 :: int)::ii) :: 36 Word.word)) ((concat_vec offset (vec_of_bits [B0,B0] :: 2 Word.word) :: 28 Word.word)) :: 64 Word.word)) \<then> - (read_reg PC_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 bits) . wGPR (vec_of_bits [B1,B1,B1,B1,B1] :: 5 Word.word) ((add_vec_int w__1 (( 8 :: int)::ii) :: 64 Word.word)))))" @@ -6708,13 +6712,13 @@ definition execute_J :: "(26)Word.word \<Rightarrow>((register_value),(unit),(e (*val execute_HCF : unit -> unit*) definition execute_HCF :: " unit \<Rightarrow> unit " where - " execute_HCF g__21 = ( () )" + " execute_HCF _ = ( () )" (*val execute_ERET : unit -> M unit*) definition execute_ERET :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_ERET g__26 = ( + " execute_ERET _ = ( (((checkCP0Access () \<then> ERETHook () ) \<then> write_reg CP0LLBit_ref (vec_of_bits [B0] :: 1 Word.word)) \<then> @@ -6871,9 +6875,9 @@ definition execute_DIVU :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 64 Word.word)))))))))) then (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . return (w__0, w__1))) else (let si = (Word.uint ((subrange_vec_dec rsVal (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))) in @@ -6897,9 +6901,9 @@ definition execute_DIV :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 64 Word.word)))))))))) then (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__0 :: 32 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> (w__1 :: 32 Word.word) . return (w__0, w__1))) else (let si = (Word.sint ((subrange_vec_dec rsVal (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word))) in @@ -6922,9 +6926,9 @@ definition execute_DDIVU :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarro (let rtVal = (Word.uint w__1) in (if (((rtVal = (( 0 :: int)::ii)))) then (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . return (w__2, w__3))) else (let qi = (hardware_quot rsVal rtVal) in @@ -6944,9 +6948,9 @@ definition execute_DDIV :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow (let rtVal = (Word.sint w__1) in (if (((rtVal = (( 0 :: int)::ii)))) then (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . return (w__2, w__3))) else (let qi = (hardware_quot rsVal rtVal) in @@ -7469,7 +7473,7 @@ definition execute_CSC :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> (*val execute_CReturn : unit -> M unit*) definition execute_CReturn :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_CReturn g__27 = ( checkCP2usable () \<then> raise_c2_exception_noreg CapEx_ReturnTrap )" + " execute_CReturn _ = ( checkCP2usable () \<then> raise_c2_exception_noreg CapEx_ReturnTrap )" (*val execute_CReadHwr : mword ty5 -> mword ty5 -> M unit*) @@ -7502,28 +7506,28 @@ definition execute_CReadHwr :: "(5)Word.word \<Rightarrow>(5)Word.word \<Righta else (let p00 = (Word.uint sel) in (if (((p00 = (( 0 :: int)::ii)))) then - (read_reg DDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__13 :: 257 Word.word) . + (read_reg DDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__13 :: CapReg) . return ((capRegToCapStruct w__13))) else if (((p00 = (( 1 :: int)::ii)))) then - (read_reg CTLSU_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__14 :: 257 Word.word) . + (read_reg CTLSU_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__14 :: CapReg) . return ((capRegToCapStruct w__14))) else if (((p00 = (( 8 :: int)::ii)))) then - (read_reg CTLSP_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__15 :: 257 Word.word) . + (read_reg CTLSP_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__15 :: CapReg) . return ((capRegToCapStruct w__15))) else if (((p00 = (( 22 :: int)::ii)))) then - (read_reg KR1C_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__16 :: 257 Word.word) . + (read_reg KR1C_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__16 :: CapReg) . return ((capRegToCapStruct w__16))) else if (((p00 = (( 23 :: int)::ii)))) then - (read_reg KR2C_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__17 :: 257 Word.word) . + (read_reg KR2C_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__17 :: CapReg) . return ((capRegToCapStruct w__17))) else if (((p00 = (( 29 :: int)::ii)))) then - (read_reg KCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__18 :: 257 Word.word) . + (read_reg KCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__18 :: CapReg) . return ((capRegToCapStruct w__18))) else if (((p00 = (( 30 :: int)::ii)))) then - (read_reg KDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__19 :: 257 Word.word) . + (read_reg KDC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__19 :: CapReg) . return ((capRegToCapStruct w__19))) else if (((p00 = (( 31 :: int)::ii)))) then - (read_reg EPCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__20 :: 257 Word.word) . + (read_reg EPCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__20 :: CapReg) . return ((capRegToCapStruct w__20))) else assert_exp False (''should be unreachable code'') \<then> undefined_CapStruct () ) \<bind> (\<lambda> (capVal :: CapStruct) . @@ -7810,9 +7814,9 @@ definition execute_CJALR :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarro else if (((((cb_ptr mod (( 4 :: int)::ii))) \<noteq> (( 0 :: int)::ii)))) then SignalException AdEL else (if link then - (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__3 :: 257 Word.word) . + (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__3 :: CapReg) . (let pcc = (capRegToCapStruct w__3) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 bits) . (let (success, linkCap) = (setCapOffset pcc ((add_vec_int w__4 (( 8 :: int)::ii) :: 64 Word.word))) in if success then writeCapReg cd1 linkCap else assert_exp False (''''))))) @@ -7938,7 +7942,7 @@ definition execute_CGetPCCSetOffset :: "(5)Word.word \<Rightarrow>(5)Word.word register_inaccessible cd1) \<bind> (\<lambda> (w__0 :: bool) . if w__0 then raise_c2_exception CapEx_AccessSystemRegsViolation cd1 else - (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__1 :: 257 Word.word) . + (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__1 :: CapReg) . (let pcc = (capRegToCapStruct w__1) in (rGPR rs :: ( 64 Word.word) M) \<bind> (\<lambda> rs_val . (let (success, newPCC) = (setCapOffset pcc rs_val) in @@ -7954,9 +7958,9 @@ definition execute_CGetPCC :: "(5)Word.word \<Rightarrow>((register_value),(uni register_inaccessible cd1) \<bind> (\<lambda> (w__0 :: bool) . if w__0 then raise_c2_exception CapEx_AccessSystemRegsViolation cd1 else - (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__1 :: 257 Word.word) . + (read_reg PCC_ref :: ( 257 Word.word) M) \<bind> (\<lambda> (w__1 :: CapReg) . (let pcc = (capRegToCapStruct w__1) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . (let (success, pcc2) = (setCapOffset pcc w__2) in assert_exp success ('''') \<then> writeCapReg cd1 pcc2))))))" @@ -8312,7 +8316,7 @@ definition execute_CBZ :: "(5)Word.word \<Rightarrow>(16)Word.word \<Rightarrow ((concat_vec imm (vec_of_bits [B0,B0] :: 2 Word.word) :: 18 Word.word)) :: 64 Word.word)) (( 4 :: int)::ii) :: 64 Word.word)) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . execute_branch ((add_vec w__2 offset :: 64 Word.word)))) else return () )))" @@ -8336,7 +8340,7 @@ definition execute_CBX :: "(5)Word.word \<Rightarrow>(16)Word.word \<Rightarrow ((concat_vec imm (vec_of_bits [B0,B0] :: 2 Word.word) :: 18 Word.word)) :: 64 Word.word)) (( 4 :: int)::ii) :: 64 Word.word)) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . execute_branch ((add_vec w__2 offset :: 64 Word.word)))) else return () )))" @@ -8378,7 +8382,7 @@ definition execute_C2Dump :: "(5)Word.word \<Rightarrow> unit " where (*val execute_BREAK : unit -> M unit*) definition execute_BREAK :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where - " execute_BREAK g__18 = ( SignalException Bp )" + " execute_BREAK _ = ( SignalException Bp )" (*val execute_BEQ : mword ty5 -> mword ty5 -> mword ty16 -> bool -> bool -> M unit*) @@ -8397,10 +8401,10 @@ definition execute_BEQ :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> ((concat_vec imm (vec_of_bits [B0,B0] :: 2 Word.word) :: 18 Word.word)) :: 64 Word.word)) (( 4 :: int)::ii) :: 64 Word.word)) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . execute_branch ((add_vec w__2 offset :: 64 Word.word)))) else if likely then - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 bits) . write_reg nextPC_ref ((add_vec_int w__3 (( 8 :: int)::ii) :: 64 Word.word))) else return () )))" @@ -8421,10 +8425,10 @@ definition execute_BCMPZ :: "(5)Word.word \<Rightarrow>(16)Word.word \<Rightarr ((concat_vec imm (vec_of_bits [B0,B0] :: 2 Word.word) :: 18 Word.word)) :: 64 Word.word)) (( 4 :: int)::ii) :: 64 Word.word)) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 bits) . execute_branch ((add_vec w__1 offset :: 64 Word.word)))) else if likely then - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 bits) . write_reg nextPC_ref ((add_vec_int w__2 (( 8 :: int)::ii) :: 64 Word.word))) else return () ) \<then> (if link then wGPR (vec_of_bits [B1,B1,B1,B1,B1] :: 5 Word.word) linkVal @@ -8597,9 +8601,9 @@ fun execute :: " ast \<Rightarrow>((register_value),(unit),(exception))monad " |" execute (JALR (rs,rd)) = ( execute_JALR rs rd )" |" execute (BEQ (rs,rd,imm,ne,likely)) = ( execute_BEQ rs rd imm ne likely )" |" execute (BCMPZ (rs,imm,cmp,link,likely)) = ( execute_BCMPZ rs imm cmp link likely )" -|" execute (SYSCALL (g__17)) = ( execute_SYSCALL g__17 )" -|" execute (BREAK (g__18)) = ( execute_BREAK g__18 )" -|" execute (WAIT (g__19)) = ( execute_WAIT g__19 )" +|" execute (SYSCALL (arg0)) = ( execute_SYSCALL arg0 )" +|" execute (BREAK (arg0)) = ( execute_BREAK arg0 )" +|" execute (WAIT (arg0)) = ( execute_WAIT arg0 )" |" execute (TRAPREG (rs,rt,cmp)) = ( execute_TRAPREG rs rt cmp )" |" execute (TRAPIMM (rs,imm,cmp)) = ( execute_TRAPIMM rs imm cmp )" |" execute (Load (width,sign,linked,base,rt,offset)) = ( execute_Load width sign linked base rt offset )" @@ -8613,16 +8617,16 @@ fun execute :: " ast \<Rightarrow>((register_value),(unit),(exception))monad " |" execute (SDL (base,rt,offset)) = ( execute_SDL base rt offset )" |" execute (SDR (base,rt,offset)) = ( execute_SDR base rt offset )" |" execute (CACHE (base,op1,imm)) = ( execute_CACHE base op1 imm )" -|" execute (SYNC (g__20)) = ( execute_SYNC g__20 )" +|" execute (SYNC (arg0)) = ( execute_SYNC arg0 )" |" execute (MFC0 (rt,rd,sel,double)) = ( execute_MFC0 rt rd sel double )" -|" execute (HCF (g__21)) = ( return ((execute_HCF g__21)))" +|" execute (HCF (arg0)) = ( return ((execute_HCF arg0)))" |" execute (MTC0 (rt,rd,sel,double)) = ( execute_MTC0 rt rd sel double )" -|" execute (TLBWI (g__22)) = ( execute_TLBWI g__22 )" -|" execute (TLBWR (g__23)) = ( execute_TLBWR g__23 )" -|" execute (TLBR (g__24)) = ( execute_TLBR g__24 )" -|" execute (TLBP (g__25)) = ( execute_TLBP g__25 )" +|" execute (TLBWI (arg0)) = ( execute_TLBWI arg0 )" +|" execute (TLBWR (arg0)) = ( execute_TLBWR arg0 )" +|" execute (TLBR (arg0)) = ( execute_TLBR arg0 )" +|" execute (TLBP (arg0)) = ( execute_TLBP arg0 )" |" execute (RDHWR (rt,rd)) = ( execute_RDHWR rt rd )" -|" execute (ERET (g__26)) = ( execute_ERET g__26 )" +|" execute (ERET (arg0)) = ( execute_ERET arg0 )" |" execute (CGetPerm (rd,cb)) = ( execute_CGetPerm rd cb )" |" execute (CGetType (rd,cb)) = ( execute_CGetType rd cb )" |" execute (CGetBase (rd,cb)) = ( execute_CGetBase rd cb )" @@ -8660,7 +8664,7 @@ fun execute :: " ast \<Rightarrow>((register_value),(unit),(exception))monad " |" execute (CCSeal (cd1,cs,ct)) = ( execute_CCSeal cd1 cs ct )" |" execute (CUnseal (cd1,cs,ct)) = ( execute_CUnseal cd1 cs ct )" |" execute (CCall (cs,cb,b__151)) = ( execute_CCall cs cb b__151 )" -|" execute (CReturn (g__27)) = ( execute_CReturn g__27 )" +|" execute (CReturn (arg0)) = ( execute_CReturn arg0 )" |" execute (CBX (cb,imm,notset)) = ( execute_CBX cb imm notset )" |" execute (CBZ (cb,imm,notzero)) = ( execute_CBZ cb imm notzero )" |" execute (CJALR (cd1,cb,link)) = ( execute_CJALR cd1 cb link )" @@ -8671,7 +8675,7 @@ fun execute :: " ast \<Rightarrow>((register_value),(unit),(exception))monad " |" execute (CSC (cs,cb,rt,rd,offset,conditional)) = ( execute_CSC cs cb rt rd offset conditional )" |" execute (CLC (cd1,cb,rt,offset,linked)) = ( execute_CLC cd1 cb rt offset linked )" |" execute (C2Dump (rt)) = ( return ((execute_C2Dump rt)))" -|" execute (RI (g__28)) = ( execute_RI g__28 )" +|" execute (RI (arg0)) = ( execute_RI arg0 )" (*val supported_instructions : ast -> maybe ast*) @@ -8695,16 +8699,16 @@ definition fetch_and_execute :: " unit \<Rightarrow>((register_value),(bool),(e (read_reg branchPending_ref :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__1 :: 1 bits) . ((write_reg inBranchDelay_ref w__1 \<then> write_reg branchPending_ref (vec_of_bits [B0] :: 1 Word.word)) \<then> - (read_reg inBranchDelay_ref :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__2 :: 1 Word.word) . + (read_reg inBranchDelay_ref :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__2 :: 1 bits) . (if ((bits_to_bool w__2)) then (read_reg delayedPC_ref :: ( 64 Word.word) M) else - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 bits) . return ((add_vec_int w__4 (( 4 :: int)::ii) :: 64 Word.word)))) \<bind> (\<lambda> (w__5 :: 64 Word.word) . ((write_reg nextPC_ref w__5 \<then> cp2_next_pc () ) \<then> read_reg instCount_ref) \<bind> (\<lambda> (w__6 :: ii) . (write_reg instCount_ref ((w__6 + (( 1 :: int)::ii))) \<then> - (read_reg UART_WRITTEN_ref :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__7 :: 1 Word.word) . + (read_reg UART_WRITTEN_ref :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__7 :: 1 bits) . (((if ((bits_to_bool w__7)) then (read_reg UART_WDATA_ref :: ( 8 Word.word) M) \<bind> (\<lambda> (w__8 :: 8 bits) . (let (_ :: unit) = (putchar ((Word.uint w__8))) in @@ -8713,7 +8717,7 @@ definition fetch_and_execute :: " unit \<Rightarrow>((register_value),(bool),(e skip () ) \<then> skip () ) \<then> ((let loop_again = True in - try_catch ((read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__9 :: 64 Word.word) . + try_catch ((read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__9 :: 64 bits) . (TranslatePC w__9 :: ( 64 Word.word) M) \<bind> (\<lambda> pc_pa . (MEMr_wrapper (( 32 :: int)::ii) pc_pa (( 4 :: int)::ii) :: ( 32 Word.word) M) \<bind> (\<lambda> instr . (let instr_ast = (decode instr) in @@ -8750,9 +8754,9 @@ definition dump_mips_state :: " unit \<Rightarrow>((register_value),(unit),(exc 64 Word.word) . return ((let _ = (print_endline - (((op@) (''DEBUG MIPS REG '') - (((op@) ((string_of_int - instance_Show_Show_Num_integer_dict idx)) (((op@) ('' '') ((string_of_bits w__1))))))))) in + (((@) (''DEBUG MIPS REG '') + (((@) ((string_of_int + instance_Show_Show_Num_integer_dict idx)) (((@) ('' '') ((string_of_bits w__1))))))))) in () ))))))))" @@ -8782,19 +8786,19 @@ definition main :: " unit \<Rightarrow>((register_value),(unit),(exception))mon definition initialize_registers :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " initialize_registers _ = ( (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . (write_reg PC_ref w__0 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (write_reg nextPC_ref w__1 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__2 :: 1 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__2 :: 1 Word.word) . (write_reg TLBProbe_ref w__2 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 6 :: int)::ii) :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__3 :: TLBIndexT) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 6 :: int)::ii) :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__3 :: 6 Word.word) . (write_reg TLBIndex_ref w__3 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 6 :: int)::ii) :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__4 :: TLBIndexT) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 6 :: int)::ii) :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__4 :: 6 Word.word) . (write_reg TLBRandom_ref w__4 \<then> undefined_TLBEntryLoReg () ) \<bind> (\<lambda> (w__5 :: TLBEntryLoReg) . (write_reg TLBEntryLo0_ref w__5 \<then> @@ -8803,10 +8807,10 @@ definition initialize_registers :: " unit \<Rightarrow>((register_value),(unit) undefined_ContextReg () ) \<bind> (\<lambda> (w__7 :: ContextReg) . (write_reg TLBContext_ref w__7 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 16 :: int)::ii) :: ( 16 Word.word) M)) \<bind> (\<lambda> (w__8 :: 16 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 16 :: int)::ii) :: ( 16 Word.word) M)) \<bind> (\<lambda> (w__8 :: 16 Word.word) . (write_reg TLBPageMask_ref w__8 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 6 :: int)::ii) :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__9 :: TLBIndexT) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 6 :: int)::ii) :: ( 6 Word.word) M)) \<bind> (\<lambda> (w__9 :: 6 Word.word) . (write_reg TLBWired_ref w__9 \<then> undefined_TLBEntryHiReg () ) \<bind> (\<lambda> (w__10 :: TLBEntryHiReg) . (write_reg TLBEntryHi_ref w__10 \<then> @@ -8941,198 +8945,198 @@ definition initialize_registers :: " unit \<Rightarrow>((register_value),(unit) undefined_TLBEntry () ) \<bind> (\<lambda> (w__75 :: TLBEntry) . (write_reg TLBEntry63_ref w__75 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__76 :: 32 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__76 :: 32 Word.word) . (write_reg CP0Compare_ref w__76 \<then> undefined_CauseReg () ) \<bind> (\<lambda> (w__77 :: CauseReg) . (write_reg CP0Cause_ref w__77 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__78 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__78 :: 64 Word.word) . (write_reg CP0EPC_ref w__78 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__79 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__79 :: 64 Word.word) . (write_reg CP0ErrorEPC_ref w__79 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__80 :: 1 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__80 :: 1 Word.word) . (write_reg CP0LLBit_ref w__80 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__81 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__81 :: 64 Word.word) . (write_reg CP0LLAddr_ref w__81 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__82 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__82 :: 64 Word.word) . (write_reg CP0BadVAddr_ref w__82 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__83 :: 32 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__83 :: 32 Word.word) . (write_reg CP0Count_ref w__83 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__84 :: 32 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 32 :: int)::ii) :: ( 32 Word.word) M)) \<bind> (\<lambda> (w__84 :: 32 Word.word) . (write_reg CP0HWREna_ref w__84 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__85 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__85 :: 64 Word.word) . (write_reg CP0UserLocal_ref w__85 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 3 :: int)::ii) :: ( 3 Word.word) M)) \<bind> (\<lambda> (w__86 :: 3 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 3 :: int)::ii) :: ( 3 Word.word) M)) \<bind> (\<lambda> (w__86 :: 3 Word.word) . (write_reg CP0ConfigK0_ref w__86 \<then> undefined_StatusReg () ) \<bind> (\<lambda> (w__87 :: StatusReg) . (write_reg CP0Status_ref w__87 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__88 :: 1 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__88 :: 1 Word.word) . (write_reg branchPending_ref w__88 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__89 :: 1 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__89 :: 1 Word.word) . (write_reg inBranchDelay_ref w__89 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__90 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__90 :: 64 Word.word) . (write_reg delayedPC_ref w__90 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__91 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__91 :: 64 Word.word) . (write_reg HI_ref w__91 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__92 :: 64 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__92 :: 64 Word.word) . (write_reg LO_ref w__92 \<then> (undefined_bitvector instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 64 :: int)::ii) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__93 :: 64 Word.word) . - (undefined_vector (( 32 :: int)::ii) w__93 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__94 :: ( 64 bits) list) . + (undefined_vector (( 32 :: int)::ii) w__93 :: ( ( 64 Word.word)list) M) \<bind> (\<lambda> (w__94 :: ( 64 Word.word) list) . (write_reg GPR_ref w__94 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 8 :: int)::ii) :: ( 8 Word.word) M)) \<bind> (\<lambda> (w__95 :: 8 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 8 :: int)::ii) :: ( 8 Word.word) M)) \<bind> (\<lambda> (w__95 :: 8 Word.word) . (write_reg UART_WDATA_ref w__95 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__96 :: 1 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__96 :: 1 Word.word) . (write_reg UART_WRITTEN_ref w__96 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 8 :: int)::ii) :: ( 8 Word.word) M)) \<bind> (\<lambda> (w__97 :: 8 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 8 :: int)::ii) :: ( 8 Word.word) M)) \<bind> (\<lambda> (w__97 :: 8 Word.word) . (write_reg UART_RDATA_ref w__97 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__98 :: 1 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__98 :: 1 Word.word) . (write_reg UART_RVALID_ref w__98 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__99 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__99 :: 257 Word.word) . (write_reg PCC_ref w__99 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__100 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__100 :: 257 Word.word) . (write_reg nextPCC_ref w__100 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__101 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__101 :: 257 Word.word) . (write_reg delayedPCC_ref w__101 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__102 :: 1 bits) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 1 :: int)::ii) :: ( 1 Word.word) M)) \<bind> (\<lambda> (w__102 :: 1 Word.word) . (write_reg inCCallDelay_ref w__102 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__103 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__103 :: 257 Word.word) . (write_reg DDC_ref w__103 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__104 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__104 :: 257 Word.word) . (write_reg C01_ref w__104 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__105 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__105 :: 257 Word.word) . (write_reg C02_ref w__105 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__106 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__106 :: 257 Word.word) . (write_reg C03_ref w__106 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__107 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__107 :: 257 Word.word) . (write_reg C04_ref w__107 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__108 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__108 :: 257 Word.word) . (write_reg C05_ref w__108 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__109 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__109 :: 257 Word.word) . (write_reg C06_ref w__109 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__110 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__110 :: 257 Word.word) . (write_reg C07_ref w__110 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__111 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__111 :: 257 Word.word) . (write_reg C08_ref w__111 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__112 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__112 :: 257 Word.word) . (write_reg C09_ref w__112 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__113 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__113 :: 257 Word.word) . (write_reg C10_ref w__113 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__114 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__114 :: 257 Word.word) . (write_reg C11_ref w__114 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__115 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__115 :: 257 Word.word) . (write_reg C12_ref w__115 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__116 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__116 :: 257 Word.word) . (write_reg C13_ref w__116 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__117 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__117 :: 257 Word.word) . (write_reg C14_ref w__117 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__118 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__118 :: 257 Word.word) . (write_reg C15_ref w__118 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__119 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__119 :: 257 Word.word) . (write_reg C16_ref w__119 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__120 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__120 :: 257 Word.word) . (write_reg C17_ref w__120 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__121 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__121 :: 257 Word.word) . (write_reg C18_ref w__121 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__122 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__122 :: 257 Word.word) . (write_reg C19_ref w__122 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__123 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__123 :: 257 Word.word) . (write_reg C20_ref w__123 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__124 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__124 :: 257 Word.word) . (write_reg C21_ref w__124 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__125 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__125 :: 257 Word.word) . (write_reg C22_ref w__125 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__126 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__126 :: 257 Word.word) . (write_reg C23_ref w__126 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__127 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__127 :: 257 Word.word) . (write_reg C24_ref w__127 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__128 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__128 :: 257 Word.word) . (write_reg C25_ref w__128 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__129 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__129 :: 257 Word.word) . (write_reg C26_ref w__129 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__130 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__130 :: 257 Word.word) . (write_reg C27_ref w__130 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__131 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__131 :: 257 Word.word) . (write_reg C28_ref w__131 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__132 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__132 :: 257 Word.word) . (write_reg C29_ref w__132 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__133 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__133 :: 257 Word.word) . (write_reg C30_ref w__133 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__134 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__134 :: 257 Word.word) . (write_reg C31_ref w__134 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__135 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__135 :: 257 Word.word) . (write_reg CTLSU_ref w__135 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__136 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__136 :: 257 Word.word) . (write_reg CTLSP_ref w__136 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__137 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__137 :: 257 Word.word) . (write_reg KR1C_ref w__137 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__138 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__138 :: 257 Word.word) . (write_reg KR2C_ref w__138 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__139 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__139 :: 257 Word.word) . (write_reg KCC_ref w__139 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__140 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__140 :: 257 Word.word) . (write_reg KDC_ref w__140 \<then> (undefined_bitvector - instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__141 :: CapReg) . + instance_Sail2_values_Bitvector_Machine_word_mword_dict (( 257 :: int)::ii) :: ( 257 Word.word) M)) \<bind> (\<lambda> (w__141 :: 257 Word.word) . (write_reg EPCC_ref w__141 \<then> undefined_CapCauseReg () ) \<bind> (\<lambda> (w__142 :: CapCauseReg) . (write_reg CapCause_ref w__142 \<then> diff --git a/snapshots/isabelle/cheri/Cheri_types.thy b/snapshots/isabelle/cheri/Cheri_types.thy index 3f439184..6446a60e 100644 --- a/snapshots/isabelle/cheri/Cheri_types.thy +++ b/snapshots/isabelle/cheri/Cheri_types.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from cheri_types.lem.\<close> +chapter \<open>Generated by Lem from \<open>cheri_types.lem\<close>.\<close> theory "Cheri_types" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_string" - "Sail2_operators_mwords" - "Sail2_prompt_monad" - "Sail2_prompt" +imports + Main + "LEM.Lem_pervasives_extra" + "Sail.Sail2_instr_kinds" + "Sail.Sail2_values" + "Sail.Sail2_string" + "Sail.Sail2_operators_mwords" + "Sail.Sail2_prompt_monad" + "Sail.Sail2_prompt" begin diff --git a/snapshots/isabelle/cheri/Mips_extras.thy b/snapshots/isabelle/cheri/Mips_extras.thy index 387b5ea7..6cabcef9 100644 --- a/snapshots/isabelle/cheri/Mips_extras.thy +++ b/snapshots/isabelle/cheri/Mips_extras.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from /tmp/sail/mips/mips_extras.lem.\<close> +chapter \<open>Generated by Lem from \<open>/auto/homes/tb592/REMS/sail/mips/mips_extras.lem\<close>.\<close> theory "Mips_extras" -imports - Main - "Lem_pervasives" - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_prompt_monad" - "Sail2_prompt" - "Sail2_operators" +imports + Main + "LEM.Lem_pervasives" + "LEM.Lem_pervasives_extra" + "Sail.Sail2_instr_kinds" + "Sail.Sail2_values" + "Sail.Sail2_prompt_monad" + "Sail.Sail2_prompt" + "Sail.Sail2_operators" begin diff --git a/snapshots/isabelle/lib/lem/Lem.thy b/snapshots/isabelle/lib/lem/Lem.thy index c6a2a883..ae2c2f2d 100644 --- a/snapshots/isabelle/lib/lem/Lem.thy +++ b/snapshots/isabelle/lib/lem/Lem.thy @@ -57,7 +57,7 @@ theory "Lem" imports LemExtraDefs - "~~/src/HOL/Word/Word" + "HOL-Word.Word" begin type_synonym numeral = nat diff --git a/snapshots/isabelle/lib/lem/LemExtraDefs.thy b/snapshots/isabelle/lib/lem/LemExtraDefs.thy index c14a669f..e43ff663 100644 --- a/snapshots/isabelle/lib/lem/LemExtraDefs.thy +++ b/snapshots/isabelle/lib/lem/LemExtraDefs.thy @@ -56,9 +56,9 @@ chapter \<open>Auxiliary Definitions needed by Lem\<close> theory "LemExtraDefs" imports - Main - "~~/src/HOL/Library/Permutation" - "~~/src/HOL/Library/While_Combinator" + Main + "HOL-Library.Permutation" + "HOL-Library.While_Combinator" begin subsection \<open>General\<close> @@ -109,7 +109,7 @@ by (induct l) auto lemma sorted_map_suc : "sorted l \<Longrightarrow> sorted (map Suc l)" -by (induct l) (simp_all add: sorted_Cons) +by (induct l) (simp_all) lemma sorted_find_indices : "sorted (find_indices P xs)" @@ -119,7 +119,7 @@ next case (Cons x xs) from sorted_map_suc[OF this] show ?case - by (simp add: sorted_Cons) + by (simp) qed lemma find_indices_set [simp] : @@ -159,7 +159,7 @@ next from sorted_find_indices[of P xs] find_indices_eq have "sorted (i # il)" by simp - hence i_leq: "\<And>i'. i' \<in> set (i # il) \<Longrightarrow> i \<le> i'" unfolding sorted_Cons by auto + hence i_leq: "\<And>i'. i' \<in> set (i # il) \<Longrightarrow> i \<le> i'" by auto from find_indices_set[of P xs, unfolded find_indices_eq] have set_i_il_eq:"\<And>i'. i' \<in> set (i # il) = (i' < length xs \<and> P (xs ! i'))" @@ -342,15 +342,15 @@ fun sorted_by :: "('a \<Rightarrow> 'a \<Rightarrow> bool)\<Rightarrow> 'a list | "sorted_by cmp (x1 # x2 # xs) = ((cmp x1 x2) \<and> sorted_by cmp (x2 # xs))" lemma sorted_by_lesseq [simp] : - "sorted_by ((op \<le>) :: ('a::{linorder}) => 'a => bool) = sorted" + "sorted_by ((\<le>) :: ('a::{linorder}) => 'a => bool) = sorted" proof (rule ext) fix l :: "'a list" - show "sorted_by (op \<le>) l = sorted l" + show "sorted_by (\<le>) l = sorted l" proof (induct l) case Nil thus ?case by simp next case (Cons x xs) - thus ?case by (cases xs) (simp_all) + thus ?case by (cases xs) (simp_all del: sorted.simps(2) add: sorted2_simps) qed qed @@ -688,7 +688,7 @@ next thus ?thesis by blast next case False - with xyl_in have "xyl \<in> op # y ` insert_in_list_at_arbitrary_pos x l" by simp + with xyl_in have "xyl \<in> (#) y ` insert_in_list_at_arbitrary_pos x l" by simp with ind_hyp obtain l1 l2 where "l = l1 @ l2 \<and> xyl = y # l1 @ x # l2" by (auto simp add: image_def Bex_def) hence "y # l = (y # l1) @ l2 \<and> xyl = (y # l1) @ [x] @ l2" by simp @@ -857,16 +857,12 @@ subsection \<open>sorting\<close> subsection \<open>Strings\<close> -lemma explode_str_simp [simp] : - "String.explode (STR l) = l" -by (metis STR_inverse UNIV_I) - declare String.literal.explode_inverse [simp] subsection \<open>num to string conversions\<close> definition nat_list_to_string :: "nat list \<Rightarrow> string" where - "nat_list_to_string nl = map char_of_nat nl" + "nat_list_to_string nl = map char_of nl" definition is_digit where "is_digit (n::nat) = (n < 10)" diff --git a/snapshots/isabelle/lib/lem/Lem_assert_extra.thy b/snapshots/isabelle/lib/lem/Lem_assert_extra.thy index b56e5a19..9dedb07a 100644 --- a/snapshots/isabelle/lib/lem/Lem_assert_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_assert_extra.thy @@ -1,10 +1,10 @@ -chapter \<open>Generated by Lem from assert_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>assert_extra.lem\<close>.\<close> theory "Lem_assert_extra" -imports - Main - "Lem" +imports + Main + "Lem" begin diff --git a/snapshots/isabelle/lib/lem/Lem_basic_classes.thy b/snapshots/isabelle/lib/lem/Lem_basic_classes.thy index c23da22b..5719357c 100644 --- a/snapshots/isabelle/lib/lem/Lem_basic_classes.thy +++ b/snapshots/isabelle/lib/lem/Lem_basic_classes.thy @@ -1,10 +1,10 @@ -chapter \<open>Generated by Lem from basic_classes.lem.\<close> +chapter \<open>Generated by Lem from \<open>basic_classes.lem\<close>.\<close> theory "Lem_basic_classes" -imports - Main - "Lem_bool" +imports + Main + "Lem_bool" begin diff --git a/snapshots/isabelle/lib/lem/Lem_bool.thy b/snapshots/isabelle/lib/lem/Lem_bool.thy index 75142160..960286a3 100644 --- a/snapshots/isabelle/lib/lem/Lem_bool.thy +++ b/snapshots/isabelle/lib/lem/Lem_bool.thy @@ -1,9 +1,9 @@ -chapter \<open>Generated by Lem from bool.lem.\<close> +chapter \<open>Generated by Lem from \<open>bool.lem\<close>.\<close> theory "Lem_bool" -imports - Main +imports + Main begin diff --git a/snapshots/isabelle/lib/lem/Lem_either.thy b/snapshots/isabelle/lib/lem/Lem_either.thy index e181f823..24fe6d67 100644 --- a/snapshots/isabelle/lib/lem/Lem_either.thy +++ b/snapshots/isabelle/lib/lem/Lem_either.thy @@ -1,13 +1,13 @@ -chapter \<open>Generated by Lem from either.lem.\<close> +chapter \<open>Generated by Lem from \<open>either.lem\<close>.\<close> theory "Lem_either" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_list" - "Lem_tuple" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_list" + "Lem_tuple" begin diff --git a/snapshots/isabelle/lib/lem/Lem_function.thy b/snapshots/isabelle/lib/lem/Lem_function.thy index 29c1fb04..b8ce9ae7 100644 --- a/snapshots/isabelle/lib/lem/Lem_function.thy +++ b/snapshots/isabelle/lib/lem/Lem_function.thy @@ -1,11 +1,11 @@ -chapter \<open>Generated by Lem from function.lem.\<close> +chapter \<open>Generated by Lem from \<open>function.lem\<close>.\<close> theory "Lem_function" -imports - Main - "Lem_bool" - "Lem_basic_classes" +imports + Main + "Lem_bool" + "Lem_basic_classes" begin diff --git a/snapshots/isabelle/lib/lem/Lem_function_extra.thy b/snapshots/isabelle/lib/lem/Lem_function_extra.thy index f742e1e6..7ed34eeb 100644 --- a/snapshots/isabelle/lib/lem/Lem_function_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_function_extra.thy @@ -1,15 +1,15 @@ -chapter \<open>Generated by Lem from function_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>function_extra.lem\<close>.\<close> theory "Lem_function_extra" -imports - Main - "Lem_maybe" - "Lem_bool" - "Lem_basic_classes" - "Lem_num" - "Lem_function" - "Lem" +imports + Main + "Lem_maybe" + "Lem_bool" + "Lem_basic_classes" + "Lem_num" + "Lem_function" + "Lem" begin diff --git a/snapshots/isabelle/lib/lem/Lem_list.thy b/snapshots/isabelle/lib/lem/Lem_list.thy index 3bdef057..1e49be6e 100644 --- a/snapshots/isabelle/lib/lem/Lem_list.thy +++ b/snapshots/isabelle/lib/lem/Lem_list.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from list.lem.\<close> +chapter \<open>Generated by Lem from \<open>list.lem\<close>.\<close> theory "Lem_list" -imports - Main - "Lem_bool" - "Lem_maybe" - "Lem_basic_classes" - "Lem_function" - "Lem_tuple" - "Lem_num" - "Lem" +imports + Main + "Lem_bool" + "Lem_maybe" + "Lem_basic_classes" + "Lem_function" + "Lem_tuple" + "Lem_num" + "Lem" begin diff --git a/snapshots/isabelle/lib/lem/Lem_list_extra.thy b/snapshots/isabelle/lib/lem/Lem_list_extra.thy index 9caf32fc..43806c73 100644 --- a/snapshots/isabelle/lib/lem/Lem_list_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_list_extra.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from list_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>list_extra.lem\<close>.\<close> theory "Lem_list_extra" -imports - Main - "Lem_bool" - "Lem_maybe" - "Lem_basic_classes" - "Lem_tuple" - "Lem_num" - "Lem_list" - "Lem_assert_extra" +imports + Main + "Lem_bool" + "Lem_maybe" + "Lem_basic_classes" + "Lem_tuple" + "Lem_num" + "Lem_list" + "Lem_assert_extra" begin diff --git a/snapshots/isabelle/lib/lem/Lem_machine_word.thy b/snapshots/isabelle/lib/lem/Lem_machine_word.thy index 85323bb2..d05e3876 100644 --- a/snapshots/isabelle/lib/lem/Lem_machine_word.thy +++ b/snapshots/isabelle/lib/lem/Lem_machine_word.thy @@ -1,15 +1,15 @@ -chapter \<open>Generated by Lem from machine_word.lem.\<close> +chapter \<open>Generated by Lem from \<open>machine_word.lem\<close>.\<close> theory "Lem_machine_word" -imports - Main - "Lem_bool" - "Lem_num" - "Lem_basic_classes" - "Lem_show" - "Lem_function" - "~~/src/HOL/Word/Word" +imports + Main + "Lem_bool" + "Lem_num" + "Lem_basic_classes" + "Lem_show" + "Lem_function" + "HOL-Word.Word" begin @@ -17,7 +17,7 @@ begin (*open import Bool Num Basic_classes Show Function*) -(*open import {isabelle} `~~/src/HOL/Word/Word`*) +(*open import {isabelle} `HOL-Word.Word`*) (*open import {hol} `wordsTheory` `wordsLib` `bitstringTheory` `integer_wordTheory`*) (*type mword 'a*) diff --git a/snapshots/isabelle/lib/lem/Lem_map.thy b/snapshots/isabelle/lib/lem/Lem_map.thy index fbaed71a..ada694f3 100644 --- a/snapshots/isabelle/lib/lem/Lem_map.thy +++ b/snapshots/isabelle/lib/lem/Lem_map.thy @@ -1,17 +1,17 @@ -chapter \<open>Generated by Lem from map.lem.\<close> +chapter \<open>Generated by Lem from \<open>map.lem\<close>.\<close> theory "Lem_map" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_function" - "Lem_maybe" - "Lem_list" - "Lem_tuple" - "Lem_set" - "Lem_num" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_function" + "Lem_maybe" + "Lem_list" + "Lem_tuple" + "Lem_set" + "Lem_num" begin diff --git a/snapshots/isabelle/lib/lem/Lem_map_extra.thy b/snapshots/isabelle/lib/lem/Lem_map_extra.thy index 4117fe81..1ed4b531 100644 --- a/snapshots/isabelle/lib/lem/Lem_map_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_map_extra.thy @@ -1,18 +1,18 @@ -chapter \<open>Generated by Lem from map_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>map_extra.lem\<close>.\<close> theory "Lem_map_extra" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_function" - "Lem_assert_extra" - "Lem_maybe" - "Lem_list" - "Lem_num" - "Lem_set" - "Lem_map" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_function" + "Lem_assert_extra" + "Lem_maybe" + "Lem_list" + "Lem_num" + "Lem_set" + "Lem_map" begin diff --git a/snapshots/isabelle/lib/lem/Lem_maybe.thy b/snapshots/isabelle/lib/lem/Lem_maybe.thy index da0bde92..6ee955c2 100644 --- a/snapshots/isabelle/lib/lem/Lem_maybe.thy +++ b/snapshots/isabelle/lib/lem/Lem_maybe.thy @@ -1,12 +1,12 @@ -chapter \<open>Generated by Lem from maybe.lem.\<close> +chapter \<open>Generated by Lem from \<open>maybe.lem\<close>.\<close> theory "Lem_maybe" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_function" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_function" begin diff --git a/snapshots/isabelle/lib/lem/Lem_maybe_extra.thy b/snapshots/isabelle/lib/lem/Lem_maybe_extra.thy index 0a57814c..80ea494f 100644 --- a/snapshots/isabelle/lib/lem/Lem_maybe_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_maybe_extra.thy @@ -1,12 +1,12 @@ -chapter \<open>Generated by Lem from maybe_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>maybe_extra.lem\<close>.\<close> theory "Lem_maybe_extra" -imports - Main - "Lem_basic_classes" - "Lem_maybe" - "Lem_assert_extra" +imports + Main + "Lem_basic_classes" + "Lem_maybe" + "Lem_assert_extra" begin diff --git a/snapshots/isabelle/lib/lem/Lem_num.thy b/snapshots/isabelle/lib/lem/Lem_num.thy index ddbdd533..8efe867b 100644 --- a/snapshots/isabelle/lib/lem/Lem_num.thy +++ b/snapshots/isabelle/lib/lem/Lem_num.thy @@ -1,21 +1,20 @@ -chapter \<open>Generated by Lem from num.lem.\<close> +chapter \<open>Generated by Lem from \<open>num.lem\<close>.\<close> theory "Lem_num" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "~~/src/HOL/Word/Word" - "Real" - "~~/src/HOL/NthRoot" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "HOL-Word.Word" + "Complex_Main" begin (*open import Bool Basic_classes*) -(*open import {isabelle} `~~/src/HOL/Word/Word` `Real` `~~/src/HOL/NthRoot`*) +(*open import {isabelle} `HOL-Word.Word` `Complex_Main`*) (*open import {hol} `integerTheory` `intReduce` `wordsTheory` `wordsLib` `ratTheory` `realTheory` `intrealTheory` `transcTheory`*) (*open import {coq} `Coq.Numbers.BinNums` `Coq.ZArith.BinInt` `Coq.ZArith.Zpower` `Coq.ZArith.Zdiv` `Coq.ZArith.Zmax` `Coq.Numbers.Natural.Peano.NPeano` `Coq.QArith.Qabs` `Coq.QArith.Qminmax` `Coq.QArith.Qround` `Coq.Reals.ROrderedType` `Coq.Reals.Rbase` `Coq.Reals.Rfunctions`*) @@ -188,15 +187,15 @@ record 'a NumPred_class= definition instance_Basic_classes_Ord_nat_dict :: "(nat)Ord_class " where " instance_Basic_classes_Ord_nat_dict = ((| - compare_method = (genericCompare (op<) (op=)), + compare_method = (genericCompare (<) (=)), - isLess_method = (op<), + isLess_method = (<), - isLessEqual_method = (op \<le>), + isLessEqual_method = (\<le>), - isGreater_method = (op>), + isGreater_method = (>), - isGreaterEqual_method = (op \<ge>)|) )" + isGreaterEqual_method = (\<ge>)|) )" (*val natAdd : nat -> nat -> nat*) @@ -204,7 +203,7 @@ definition instance_Basic_classes_Ord_nat_dict :: "(nat)Ord_class " where definition instance_Num_NumAdd_nat_dict :: "(nat)NumAdd_class " where " instance_Num_NumAdd_nat_dict = ((| - numAdd_method = (op+)|) )" + numAdd_method = (+)|) )" (*val natMinus : nat -> nat -> nat*) @@ -212,7 +211,7 @@ definition instance_Num_NumAdd_nat_dict :: "(nat)NumAdd_class " where definition instance_Num_NumMinus_nat_dict :: "(nat)NumMinus_class " where " instance_Num_NumMinus_nat_dict = ((| - numMinus_method = (op-)|) )" + numMinus_method = (-)|) )" (*val natSucc : nat -> nat*) @@ -235,7 +234,7 @@ definition instance_Num_NumPred_nat_dict :: "(nat)NumPred_class " where definition instance_Num_NumMult_nat_dict :: "(nat)NumMult_class " where " instance_Num_NumMult_nat_dict = ((| - numMult_method = (op*)|) )" + numMult_method = ( * )|) )" (*val natDiv : nat -> nat -> nat*) @@ -243,13 +242,13 @@ definition instance_Num_NumMult_nat_dict :: "(nat)NumMult_class " where definition instance_Num_NumIntegerDivision_nat_dict :: "(nat)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_nat_dict = ((| - div_method = (op div)|) )" + div_method = (div)|) )" definition instance_Num_NumDivision_nat_dict :: "(nat)NumDivision_class " where " instance_Num_NumDivision_nat_dict = ((| - numDivision_method = (op div)|) )" + numDivision_method = (div)|) )" (*val natMod : nat -> nat -> nat*) @@ -257,7 +256,7 @@ definition instance_Num_NumDivision_nat_dict :: "(nat)NumDivision_class " wher definition instance_Num_NumRemainder_nat_dict :: "(nat)NumRemainder_class " where " instance_Num_NumRemainder_nat_dict = ((| - mod_method = (op mod)|) )" + mod_method = (mod)|) )" @@ -284,7 +283,7 @@ definition gen_pow :: " 'a \<Rightarrow>('a \<Rightarrow> 'a \<Rightarrow> 'a)\ definition instance_Num_NumPow_nat_dict :: "(nat)NumPow_class " where " instance_Num_NumPow_nat_dict = ((| - numPow_method = (op^)|) )" + numPow_method = (^)|) )" (*val natMin : nat -> nat -> nat*) @@ -318,15 +317,15 @@ definition instance_Basic_classes_OrdMaxMin_nat_dict :: "(nat)OrdMaxMin_class " definition instance_Basic_classes_Ord_Num_natural_dict :: "(nat)Ord_class " where " instance_Basic_classes_Ord_Num_natural_dict = ((| - compare_method = (genericCompare (op<) (op=)), + compare_method = (genericCompare (<) (=)), - isLess_method = (op<), + isLess_method = (<), - isLessEqual_method = (op \<le>), + isLessEqual_method = (\<le>), - isGreater_method = (op>), + isGreater_method = (>), - isGreaterEqual_method = (op \<ge>)|) )" + isGreaterEqual_method = (\<ge>)|) )" (*val naturalAdd : natural -> natural -> natural*) @@ -334,7 +333,7 @@ definition instance_Basic_classes_Ord_Num_natural_dict :: "(nat)Ord_class " wh definition instance_Num_NumAdd_Num_natural_dict :: "(nat)NumAdd_class " where " instance_Num_NumAdd_Num_natural_dict = ((| - numAdd_method = (op+)|) )" + numAdd_method = (+)|) )" (*val naturalMinus : natural -> natural -> natural*) @@ -342,7 +341,7 @@ definition instance_Num_NumAdd_Num_natural_dict :: "(nat)NumAdd_class " where definition instance_Num_NumMinus_Num_natural_dict :: "(nat)NumMinus_class " where " instance_Num_NumMinus_Num_natural_dict = ((| - numMinus_method = (op-)|) )" + numMinus_method = (-)|) )" (*val naturalSucc : natural -> natural*) @@ -365,7 +364,7 @@ definition instance_Num_NumPred_Num_natural_dict :: "(nat)NumPred_class " wher definition instance_Num_NumMult_Num_natural_dict :: "(nat)NumMult_class " where " instance_Num_NumMult_Num_natural_dict = ((| - numMult_method = (op*)|) )" + numMult_method = ( * )|) )" @@ -374,7 +373,7 @@ definition instance_Num_NumMult_Num_natural_dict :: "(nat)NumMult_class " wher definition instance_Num_NumPow_Num_natural_dict :: "(nat)NumPow_class " where " instance_Num_NumPow_Num_natural_dict = ((| - numPow_method = (op^)|) )" + numPow_method = (^)|) )" (*val naturalDiv : natural -> natural -> natural*) @@ -382,13 +381,13 @@ definition instance_Num_NumPow_Num_natural_dict :: "(nat)NumPow_class " where definition instance_Num_NumIntegerDivision_Num_natural_dict :: "(nat)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_natural_dict = ((| - div_method = (op div)|) )" + div_method = (div)|) )" definition instance_Num_NumDivision_Num_natural_dict :: "(nat)NumDivision_class " where " instance_Num_NumDivision_Num_natural_dict = ((| - numDivision_method = (op div)|) )" + numDivision_method = (div)|) )" (*val naturalMod : natural -> natural -> natural*) @@ -396,7 +395,7 @@ definition instance_Num_NumDivision_Num_natural_dict :: "(nat)NumDivision_class definition instance_Num_NumRemainder_Num_natural_dict :: "(nat)NumRemainder_class " where " instance_Num_NumRemainder_Num_natural_dict = ((| - mod_method = (op mod)|) )" + mod_method = (mod)|) )" (*val naturalMin : natural -> natural -> natural*) @@ -430,15 +429,15 @@ definition instance_Basic_classes_OrdMaxMin_Num_natural_dict :: "(nat)OrdMaxMin definition instance_Basic_classes_Ord_Num_int_dict :: "(int)Ord_class " where " instance_Basic_classes_Ord_Num_int_dict = ((| - compare_method = (genericCompare (op<) (op=)), + compare_method = (genericCompare (<) (=)), - isLess_method = (op<), + isLess_method = (<), - isLessEqual_method = (op \<le>), + isLessEqual_method = (\<le>), - isGreater_method = (op>), + isGreater_method = (>), - isGreaterEqual_method = (op \<ge>)|) )" + isGreaterEqual_method = (\<ge>)|) )" (*val intNegate : int -> int*) @@ -462,7 +461,7 @@ definition instance_Num_NumAbs_Num_int_dict :: "(int)NumAbs_class " where definition instance_Num_NumAdd_Num_int_dict :: "(int)NumAdd_class " where " instance_Num_NumAdd_Num_int_dict = ((| - numAdd_method = (op+)|) )" + numAdd_method = (+)|) )" (*val intMinus : int -> int -> int*) @@ -470,7 +469,7 @@ definition instance_Num_NumAdd_Num_int_dict :: "(int)NumAdd_class " where definition instance_Num_NumMinus_Num_int_dict :: "(int)NumMinus_class " where " instance_Num_NumMinus_Num_int_dict = ((| - numMinus_method = (op-)|) )" + numMinus_method = (-)|) )" (*val intSucc : int -> int*) @@ -492,7 +491,7 @@ definition instance_Num_NumPred_Num_int_dict :: "(int)NumPred_class " where definition instance_Num_NumMult_Num_int_dict :: "(int)NumMult_class " where " instance_Num_NumMult_Num_int_dict = ((| - numMult_method = (op*)|) )" + numMult_method = ( * )|) )" @@ -501,7 +500,7 @@ definition instance_Num_NumMult_Num_int_dict :: "(int)NumMult_class " where definition instance_Num_NumPow_Num_int_dict :: "(int)NumPow_class " where " instance_Num_NumPow_Num_int_dict = ((| - numPow_method = (op^)|) )" + numPow_method = (^)|) )" (*val intDiv : int -> int -> int*) @@ -509,13 +508,13 @@ definition instance_Num_NumPow_Num_int_dict :: "(int)NumPow_class " where definition instance_Num_NumIntegerDivision_Num_int_dict :: "(int)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_int_dict = ((| - div_method = (op div)|) )" + div_method = (div)|) )" definition instance_Num_NumDivision_Num_int_dict :: "(int)NumDivision_class " where " instance_Num_NumDivision_Num_int_dict = ((| - numDivision_method = (op div)|) )" + numDivision_method = (div)|) )" (*val intMod : int -> int -> int*) @@ -523,7 +522,7 @@ definition instance_Num_NumDivision_Num_int_dict :: "(int)NumDivision_class " definition instance_Num_NumRemainder_Num_int_dict :: "(int)NumRemainder_class " where " instance_Num_NumRemainder_Num_int_dict = ((| - mod_method = (op mod)|) )" + mod_method = (mod)|) )" (*val intMin : int -> int -> int*) @@ -555,7 +554,7 @@ definition instance_Basic_classes_OrdMaxMin_Num_int_dict :: "(int)OrdMaxMin_cla definition instance_Basic_classes_Ord_Num_int32_dict :: "( 32 word)Ord_class " where " instance_Basic_classes_Ord_Num_int32_dict = ((| - compare_method = (genericCompare word_sless (op=)), + compare_method = (genericCompare word_sless (=)), isLess_method = word_sless, @@ -591,7 +590,7 @@ definition instance_Num_NumAbs_Num_int32_dict :: "( 32 word)NumAbs_class " whe definition instance_Num_NumAdd_Num_int32_dict :: "( 32 word)NumAdd_class " where " instance_Num_NumAdd_Num_int32_dict = ((| - numAdd_method = (op+)|) )" + numAdd_method = (+)|) )" (*val int32Minus : int32 -> int32 -> int32*) @@ -599,7 +598,7 @@ definition instance_Num_NumAdd_Num_int32_dict :: "( 32 word)NumAdd_class " whe definition instance_Num_NumMinus_Num_int32_dict :: "( 32 word)NumMinus_class " where " instance_Num_NumMinus_Num_int32_dict = ((| - numMinus_method = (op-)|) )" + numMinus_method = (-)|) )" (*val int32Succ : int32 -> int32*) @@ -622,7 +621,7 @@ definition instance_Num_NumPred_Num_int32_dict :: "( 32 word)NumPred_class " w definition instance_Num_NumMult_Num_int32_dict :: "( 32 word)NumMult_class " where " instance_Num_NumMult_Num_int32_dict = ((| - numMult_method = (op*)|) )" + numMult_method = ( * )|) )" @@ -631,7 +630,7 @@ definition instance_Num_NumMult_Num_int32_dict :: "( 32 word)NumMult_class " w definition instance_Num_NumPow_Num_int32_dict :: "( 32 word)NumPow_class " where " instance_Num_NumPow_Num_int32_dict = ((| - numPow_method = (op^)|) )" + numPow_method = (^)|) )" (*val int32Div : int32 -> int32 -> int32*) @@ -639,13 +638,13 @@ definition instance_Num_NumPow_Num_int32_dict :: "( 32 word)NumPow_class " whe definition instance_Num_NumIntegerDivision_Num_int32_dict :: "( 32 word)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_int32_dict = ((| - div_method = (op div)|) )" + div_method = (div)|) )" definition instance_Num_NumDivision_Num_int32_dict :: "( 32 word)NumDivision_class " where " instance_Num_NumDivision_Num_int32_dict = ((| - numDivision_method = (op div)|) )" + numDivision_method = (div)|) )" (*val int32Mod : int32 -> int32 -> int32*) @@ -653,7 +652,7 @@ definition instance_Num_NumDivision_Num_int32_dict :: "( 32 word)NumDivision_cl definition instance_Num_NumRemainder_Num_int32_dict :: "( 32 word)NumRemainder_class " where " instance_Num_NumRemainder_Num_int32_dict = ((| - mod_method = (op mod)|) )" + mod_method = (mod)|) )" (*val int32Min : int32 -> int32 -> int32*) @@ -687,7 +686,7 @@ definition instance_Basic_classes_OrdMaxMin_Num_int32_dict :: "( 32 word)OrdMax definition instance_Basic_classes_Ord_Num_int64_dict :: "( 64 word)Ord_class " where " instance_Basic_classes_Ord_Num_int64_dict = ((| - compare_method = (genericCompare word_sless (op=)), + compare_method = (genericCompare word_sless (=)), isLess_method = word_sless, @@ -723,7 +722,7 @@ definition instance_Num_NumAbs_Num_int64_dict :: "( 64 word)NumAbs_class " whe definition instance_Num_NumAdd_Num_int64_dict :: "( 64 word)NumAdd_class " where " instance_Num_NumAdd_Num_int64_dict = ((| - numAdd_method = (op+)|) )" + numAdd_method = (+)|) )" (*val int64Minus : int64 -> int64 -> int64*) @@ -731,7 +730,7 @@ definition instance_Num_NumAdd_Num_int64_dict :: "( 64 word)NumAdd_class " whe definition instance_Num_NumMinus_Num_int64_dict :: "( 64 word)NumMinus_class " where " instance_Num_NumMinus_Num_int64_dict = ((| - numMinus_method = (op-)|) )" + numMinus_method = (-)|) )" (*val int64Succ : int64 -> int64*) @@ -754,7 +753,7 @@ definition instance_Num_NumPred_Num_int64_dict :: "( 64 word)NumPred_class " w definition instance_Num_NumMult_Num_int64_dict :: "( 64 word)NumMult_class " where " instance_Num_NumMult_Num_int64_dict = ((| - numMult_method = (op*)|) )" + numMult_method = ( * )|) )" @@ -763,7 +762,7 @@ definition instance_Num_NumMult_Num_int64_dict :: "( 64 word)NumMult_class " w definition instance_Num_NumPow_Num_int64_dict :: "( 64 word)NumPow_class " where " instance_Num_NumPow_Num_int64_dict = ((| - numPow_method = (op^)|) )" + numPow_method = (^)|) )" (*val int64Div : int64 -> int64 -> int64*) @@ -771,13 +770,13 @@ definition instance_Num_NumPow_Num_int64_dict :: "( 64 word)NumPow_class " whe definition instance_Num_NumIntegerDivision_Num_int64_dict :: "( 64 word)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_int64_dict = ((| - div_method = (op div)|) )" + div_method = (div)|) )" definition instance_Num_NumDivision_Num_int64_dict :: "( 64 word)NumDivision_class " where " instance_Num_NumDivision_Num_int64_dict = ((| - numDivision_method = (op div)|) )" + numDivision_method = (div)|) )" (*val int64Mod : int64 -> int64 -> int64*) @@ -785,7 +784,7 @@ definition instance_Num_NumDivision_Num_int64_dict :: "( 64 word)NumDivision_cl definition instance_Num_NumRemainder_Num_int64_dict :: "( 64 word)NumRemainder_class " where " instance_Num_NumRemainder_Num_int64_dict = ((| - mod_method = (op mod)|) )" + mod_method = (mod)|) )" (*val int64Min : int64 -> int64 -> int64*) @@ -821,15 +820,15 @@ definition instance_Basic_classes_OrdMaxMin_Num_int64_dict :: "( 64 word)OrdMax definition instance_Basic_classes_Ord_Num_integer_dict :: "(int)Ord_class " where " instance_Basic_classes_Ord_Num_integer_dict = ((| - compare_method = (genericCompare (op<) (op=)), + compare_method = (genericCompare (<) (=)), - isLess_method = (op<), + isLess_method = (<), - isLessEqual_method = (op \<le>), + isLessEqual_method = (\<le>), - isGreater_method = (op>), + isGreater_method = (>), - isGreaterEqual_method = (op \<ge>)|) )" + isGreaterEqual_method = (\<ge>)|) )" (*val integerNegate : integer -> integer*) @@ -853,7 +852,7 @@ definition instance_Num_NumAbs_Num_integer_dict :: "(int)NumAbs_class " where definition instance_Num_NumAdd_Num_integer_dict :: "(int)NumAdd_class " where " instance_Num_NumAdd_Num_integer_dict = ((| - numAdd_method = (op+)|) )" + numAdd_method = (+)|) )" (*val integerMinus : integer -> integer -> integer*) @@ -861,7 +860,7 @@ definition instance_Num_NumAdd_Num_integer_dict :: "(int)NumAdd_class " where definition instance_Num_NumMinus_Num_integer_dict :: "(int)NumMinus_class " where " instance_Num_NumMinus_Num_integer_dict = ((| - numMinus_method = (op-)|) )" + numMinus_method = (-)|) )" (*val integerSucc : integer -> integer*) @@ -883,7 +882,7 @@ definition instance_Num_NumPred_Num_integer_dict :: "(int)NumPred_class " wher definition instance_Num_NumMult_Num_integer_dict :: "(int)NumMult_class " where " instance_Num_NumMult_Num_integer_dict = ((| - numMult_method = (op*)|) )" + numMult_method = ( * )|) )" @@ -892,7 +891,7 @@ definition instance_Num_NumMult_Num_integer_dict :: "(int)NumMult_class " wher definition instance_Num_NumPow_Num_integer_dict :: "(int)NumPow_class " where " instance_Num_NumPow_Num_integer_dict = ((| - numPow_method = (op^)|) )" + numPow_method = (^)|) )" (*val integerDiv : integer -> integer -> integer*) @@ -900,13 +899,13 @@ definition instance_Num_NumPow_Num_integer_dict :: "(int)NumPow_class " where definition instance_Num_NumIntegerDivision_Num_integer_dict :: "(int)NumIntegerDivision_class " where " instance_Num_NumIntegerDivision_Num_integer_dict = ((| - div_method = (op div)|) )" + div_method = (div)|) )" definition instance_Num_NumDivision_Num_integer_dict :: "(int)NumDivision_class " where " instance_Num_NumDivision_Num_integer_dict = ((| - numDivision_method = (op div)|) )" + numDivision_method = (div)|) )" (*val integerMod : integer -> integer -> integer*) @@ -914,7 +913,7 @@ definition instance_Num_NumDivision_Num_integer_dict :: "(int)NumDivision_class definition instance_Num_NumRemainder_Num_integer_dict :: "(int)NumRemainder_class " where " instance_Num_NumRemainder_Num_integer_dict = ((| - mod_method = (op mod)|) )" + mod_method = (mod)|) )" (*val integerMin : integer -> integer -> integer*) @@ -951,15 +950,15 @@ definition instance_Basic_classes_OrdMaxMin_Num_integer_dict :: "(int)OrdMaxMin definition instance_Basic_classes_Ord_Num_rational_dict :: "(rat)Ord_class " where " instance_Basic_classes_Ord_Num_rational_dict = ((| - compare_method = (genericCompare (op<) (op=)), + compare_method = (genericCompare (<) (=)), - isLess_method = (op<), + isLess_method = (<), - isLessEqual_method = (op \<le>), + isLessEqual_method = (\<le>), - isGreater_method = (op>), + isGreater_method = (>), - isGreaterEqual_method = (op \<ge>)|) )" + isGreaterEqual_method = (\<ge>)|) )" (*val rationalAdd : rational -> rational -> rational*) @@ -967,7 +966,7 @@ definition instance_Basic_classes_Ord_Num_rational_dict :: "(rat)Ord_class " w definition instance_Num_NumAdd_Num_rational_dict :: "(rat)NumAdd_class " where " instance_Num_NumAdd_Num_rational_dict = ((| - numAdd_method = (op+)|) )" + numAdd_method = (+)|) )" (*val rationalMinus : rational -> rational -> rational*) @@ -975,7 +974,7 @@ definition instance_Num_NumAdd_Num_rational_dict :: "(rat)NumAdd_class " where definition instance_Num_NumMinus_Num_rational_dict :: "(rat)NumMinus_class " where " instance_Num_NumMinus_Num_rational_dict = ((| - numMinus_method = (op-)|) )" + numMinus_method = (-)|) )" (*val rationalNegate : rational -> rational*) @@ -1013,7 +1012,7 @@ definition instance_Num_NumPred_Num_rational_dict :: "(rat)NumPred_class " whe definition instance_Num_NumMult_Num_rational_dict :: "(rat)NumMult_class " where " instance_Num_NumMult_Num_rational_dict = ((| - numMult_method = (op*)|) )" + numMult_method = ( * )|) )" (*val rationalDiv : rational -> rational -> rational*) @@ -1021,7 +1020,7 @@ definition instance_Num_NumMult_Num_rational_dict :: "(rat)NumMult_class " whe definition instance_Num_NumDivision_Num_rational_dict :: "(rat)NumDivision_class " where " instance_Num_NumDivision_Num_rational_dict = ((| - numDivision_method = (op div)|) )" + numDivision_method = (div)|) )" (*val rationalFromFrac : int -> int -> rational*) @@ -1078,15 +1077,15 @@ definition instance_Basic_classes_OrdMaxMin_Num_rational_dict :: "(rat)OrdMaxMi definition instance_Basic_classes_Ord_Num_real_dict :: "(real)Ord_class " where " instance_Basic_classes_Ord_Num_real_dict = ((| - compare_method = (genericCompare (op<) (op=)), + compare_method = (genericCompare (<) (=)), - isLess_method = (op<), + isLess_method = (<), - isLessEqual_method = (op \<le>), + isLessEqual_method = (\<le>), - isGreater_method = (op>), + isGreater_method = (>), - isGreaterEqual_method = (op \<ge>)|) )" + isGreaterEqual_method = (\<ge>)|) )" (*val realAdd : real -> real -> real*) @@ -1094,7 +1093,7 @@ definition instance_Basic_classes_Ord_Num_real_dict :: "(real)Ord_class " wher definition instance_Num_NumAdd_Num_real_dict :: "(real)NumAdd_class " where " instance_Num_NumAdd_Num_real_dict = ((| - numAdd_method = (op+)|) )" + numAdd_method = (+)|) )" (*val realMinus : real -> real -> real*) @@ -1102,7 +1101,7 @@ definition instance_Num_NumAdd_Num_real_dict :: "(real)NumAdd_class " where definition instance_Num_NumMinus_Num_real_dict :: "(real)NumMinus_class " where " instance_Num_NumMinus_Num_real_dict = ((| - numMinus_method = (op-)|) )" + numMinus_method = (-)|) )" (*val realNegate : real -> real*) @@ -1140,7 +1139,7 @@ definition instance_Num_NumPred_Num_real_dict :: "(real)NumPred_class " where definition instance_Num_NumMult_Num_real_dict :: "(real)NumMult_class " where " instance_Num_NumMult_Num_real_dict = ((| - numMult_method = (op*)|) )" + numMult_method = ( * )|) )" (*val realDiv : real -> real -> real*) @@ -1148,7 +1147,7 @@ definition instance_Num_NumMult_Num_real_dict :: "(real)NumMult_class " where definition instance_Num_NumDivision_Num_real_dict :: "(real)NumDivision_class " where " instance_Num_NumDivision_Num_real_dict = ((| - numDivision_method = (op div)|) )" + numDivision_method = (div)|) )" (*val realFromFrac : integer -> integer -> real*) diff --git a/snapshots/isabelle/lib/lem/Lem_num_extra.thy b/snapshots/isabelle/lib/lem/Lem_num_extra.thy index 0611862e..c1edb194 100644 --- a/snapshots/isabelle/lib/lem/Lem_num_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_num_extra.thy @@ -1,11 +1,11 @@ -chapter \<open>Generated by Lem from num_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>num_extra.lem\<close>.\<close> theory "Lem_num_extra" -imports - Main - "Lem_num" - "Lem_string" +imports + Main + "Lem_num" + "Lem_string" begin diff --git a/snapshots/isabelle/lib/lem/Lem_pervasives.thy b/snapshots/isabelle/lib/lem/Lem_pervasives.thy index 37da1224..b8f1ffd3 100644 --- a/snapshots/isabelle/lib/lem/Lem_pervasives.thy +++ b/snapshots/isabelle/lib/lem/Lem_pervasives.thy @@ -1,24 +1,24 @@ -chapter \<open>Generated by Lem from pervasives.lem.\<close> +chapter \<open>Generated by Lem from \<open>pervasives.lem\<close>.\<close> theory "Lem_pervasives" -imports - Main - "Lem_basic_classes" - "Lem_bool" - "Lem_tuple" - "Lem_maybe" - "Lem_either" - "Lem_function" - "Lem_num" - "Lem_map" - "Lem_set" - "Lem_list" - "Lem_string" - "Lem_word" - "Lem_show" - "Lem_sorting" - "Lem_relation" +imports + Main + "Lem_basic_classes" + "Lem_bool" + "Lem_tuple" + "Lem_maybe" + "Lem_either" + "Lem_function" + "Lem_num" + "Lem_map" + "Lem_set" + "Lem_list" + "Lem_string" + "Lem_word" + "Lem_show" + "Lem_sorting" + "Lem_relation" begin diff --git a/snapshots/isabelle/lib/lem/Lem_pervasives_extra.thy b/snapshots/isabelle/lib/lem/Lem_pervasives_extra.thy index 0e3e5a88..ba3ba259 100644 --- a/snapshots/isabelle/lib/lem/Lem_pervasives_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_pervasives_extra.thy @@ -1,21 +1,21 @@ -chapter \<open>Generated by Lem from pervasives_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>pervasives_extra.lem\<close>.\<close> theory "Lem_pervasives_extra" -imports - Main - "Lem_pervasives" - "Lem_function_extra" - "Lem_maybe_extra" - "Lem_map_extra" - "Lem_num_extra" - "Lem_set_extra" - "Lem_set_helpers" - "Lem_list_extra" - "Lem_string_extra" - "Lem_assert_extra" - "Lem_show_extra" - "Lem_machine_word" +imports + Main + "Lem_pervasives" + "Lem_function_extra" + "Lem_maybe_extra" + "Lem_map_extra" + "Lem_num_extra" + "Lem_set_extra" + "Lem_set_helpers" + "Lem_list_extra" + "Lem_string_extra" + "Lem_assert_extra" + "Lem_show_extra" + "Lem_machine_word" begin diff --git a/snapshots/isabelle/lib/lem/Lem_relation.thy b/snapshots/isabelle/lib/lem/Lem_relation.thy index 23e7d707..778507a7 100644 --- a/snapshots/isabelle/lib/lem/Lem_relation.thy +++ b/snapshots/isabelle/lib/lem/Lem_relation.thy @@ -1,14 +1,14 @@ -chapter \<open>Generated by Lem from relation.lem.\<close> +chapter \<open>Generated by Lem from \<open>relation.lem\<close>.\<close> theory "Lem_relation" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_tuple" - "Lem_set" - "Lem_num" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_tuple" + "Lem_set" + "Lem_num" begin @@ -85,7 +85,7 @@ definition relFromPred :: " 'a set \<Rightarrow> 'b set \<Rightarrow>('a \<Righ (*val relIdOn : forall 'a. SetType 'a, Eq 'a => set 'a -> rel 'a 'a*) definition relIdOn :: " 'a set \<Rightarrow>('a*'a)set " where - " relIdOn s = ( relFromPred s s (op=))" + " relIdOn s = ( relFromPred s s (=))" (*val relId : forall 'a. SetType 'a, Eq 'a => rel 'a 'a*) diff --git a/snapshots/isabelle/lib/lem/Lem_set.thy b/snapshots/isabelle/lib/lem/Lem_set.thy index f77d4d98..ab6ee68c 100644 --- a/snapshots/isabelle/lib/lem/Lem_set.thy +++ b/snapshots/isabelle/lib/lem/Lem_set.thy @@ -1,17 +1,17 @@ -chapter \<open>Generated by Lem from set.lem.\<close> +chapter \<open>Generated by Lem from \<open>set.lem\<close>.\<close> theory "Lem_set" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_maybe" - "Lem_function" - "Lem_num" - "Lem_list" - "Lem_set_helpers" - "Lem" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_maybe" + "Lem_function" + "Lem_num" + "Lem_list" + "Lem_set_helpers" + "Lem" begin diff --git a/snapshots/isabelle/lib/lem/Lem_set_extra.thy b/snapshots/isabelle/lib/lem/Lem_set_extra.thy index 33516be7..87ccacef 100644 --- a/snapshots/isabelle/lib/lem/Lem_set_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_set_extra.thy @@ -1,17 +1,17 @@ -chapter \<open>Generated by Lem from set_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>set_extra.lem\<close>.\<close> theory "Lem_set_extra" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_maybe" - "Lem_function" - "Lem_num" - "Lem_list" - "Lem_sorting" - "Lem_set" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_maybe" + "Lem_function" + "Lem_num" + "Lem_list" + "Lem_sorting" + "Lem_set" begin diff --git a/snapshots/isabelle/lib/lem/Lem_set_helpers.thy b/snapshots/isabelle/lib/lem/Lem_set_helpers.thy index 1a2f5f50..3b8991cf 100644 --- a/snapshots/isabelle/lib/lem/Lem_set_helpers.thy +++ b/snapshots/isabelle/lib/lem/Lem_set_helpers.thy @@ -1,14 +1,14 @@ -chapter \<open>Generated by Lem from set_helpers.lem.\<close> +chapter \<open>Generated by Lem from \<open>set_helpers.lem\<close>.\<close> theory "Lem_set_helpers" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_maybe" - "Lem_function" - "Lem_num" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_maybe" + "Lem_function" + "Lem_num" begin diff --git a/snapshots/isabelle/lib/lem/Lem_show.thy b/snapshots/isabelle/lib/lem/Lem_show.thy index fced534d..b55de519 100644 --- a/snapshots/isabelle/lib/lem/Lem_show.thy +++ b/snapshots/isabelle/lib/lem/Lem_show.thy @@ -1,13 +1,13 @@ -chapter \<open>Generated by Lem from show.lem.\<close> +chapter \<open>Generated by Lem from \<open>show.lem\<close>.\<close> theory "Lem_show" -imports - Main - "Lem_string" - "Lem_maybe" - "Lem_num" - "Lem_basic_classes" +imports + Main + "Lem_string" + "Lem_maybe" + "Lem_num" + "Lem_basic_classes" begin @@ -26,7 +26,7 @@ record 'a Show_class= definition instance_Show_Show_string_dict :: "(string)Show_class " where " instance_Show_Show_string_dict = ((| - show_method = (\<lambda> s. ([(char_of_nat 34)]) @ (s @ ([(char_of_nat 34)])))|) )" + show_method = (\<lambda> s. ([(CHR 0x27)]) @ (s @ ([(CHR 0x27)])))|) )" (*val stringFromMaybe : forall 'a. ('a -> string) -> maybe 'a -> string*) diff --git a/snapshots/isabelle/lib/lem/Lem_show_extra.thy b/snapshots/isabelle/lib/lem/Lem_show_extra.thy index 25ab2570..36e23aca 100644 --- a/snapshots/isabelle/lib/lem/Lem_show_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_show_extra.thy @@ -1,18 +1,18 @@ -chapter \<open>Generated by Lem from show_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>show_extra.lem\<close>.\<close> theory "Lem_show_extra" -imports - Main - "Lem_string" - "Lem_maybe" - "Lem_num" - "Lem_basic_classes" - "Lem_set" - "Lem_relation" - "Lem_show" - "Lem_set_extra" - "Lem_string_extra" +imports + Main + "Lem_string" + "Lem_maybe" + "Lem_num" + "Lem_basic_classes" + "Lem_set" + "Lem_relation" + "Lem_show" + "Lem_set_extra" + "Lem_string_extra" begin diff --git a/snapshots/isabelle/lib/lem/Lem_sorting.thy b/snapshots/isabelle/lib/lem/Lem_sorting.thy index d42425a2..48032c8e 100644 --- a/snapshots/isabelle/lib/lem/Lem_sorting.thy +++ b/snapshots/isabelle/lib/lem/Lem_sorting.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from sorting.lem.\<close> +chapter \<open>Generated by Lem from \<open>sorting.lem\<close>.\<close> theory "Lem_sorting" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_maybe" - "Lem_list" - "Lem_num" - "Lem" - "~~/src/HOL/Library/Permutation" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_maybe" + "Lem_list" + "Lem_num" + "Lem" + "HOL-Library.Permutation" begin @@ -18,7 +18,7 @@ begin (*open import Bool Basic_classes Maybe List Num*) -(*open import {isabelle} `~~/src/HOL/Library/Permutation`*) +(*open import {isabelle} `HOL-Library.Permutation`*) (*open import {coq} `Coq.Lists.List`*) (*open import {hol} `sortingTheory` `permLib`*) (*open import {isabelle} `$LIB_DIR/Lem`*) diff --git a/snapshots/isabelle/lib/lem/Lem_string.thy b/snapshots/isabelle/lib/lem/Lem_string.thy index 9df246c4..7fd329bd 100644 --- a/snapshots/isabelle/lib/lem/Lem_string.thy +++ b/snapshots/isabelle/lib/lem/Lem_string.thy @@ -1,12 +1,12 @@ -chapter \<open>Generated by Lem from string.lem.\<close> +chapter \<open>Generated by Lem from \<open>string.lem\<close>.\<close> theory "Lem_string" -imports - Main - "Lem_bool" - "Lem_basic_classes" - "Lem_list" +imports + Main + "Lem_bool" + "Lem_basic_classes" + "Lem_list" begin diff --git a/snapshots/isabelle/lib/lem/Lem_string_extra.thy b/snapshots/isabelle/lib/lem/Lem_string_extra.thy index bd8317ba..9991308f 100644 --- a/snapshots/isabelle/lib/lem/Lem_string_extra.thy +++ b/snapshots/isabelle/lib/lem/Lem_string_extra.thy @@ -1,14 +1,14 @@ -chapter \<open>Generated by Lem from string_extra.lem.\<close> +chapter \<open>Generated by Lem from \<open>string_extra.lem\<close>.\<close> theory "Lem_string_extra" -imports - Main - "Lem_num" - "Lem_list" - "Lem_basic_classes" - "Lem_string" - "Lem_list_extra" +imports + Main + "Lem_num" + "Lem_list" + "Lem_basic_classes" + "Lem_string" + "Lem_list_extra" begin @@ -43,7 +43,7 @@ fun stringFromNatHelper :: " nat \<Rightarrow>(char)list \<Rightarrow>(char)li if n =( 0 :: nat) then acc1 else - stringFromNatHelper (n div( 10 :: nat)) (char_of_nat ((n mod( 10 :: nat)) +( 48 :: nat)) # acc1))" + stringFromNatHelper (n div( 10 :: nat)) ((%n. char_of (n::nat)) ((n mod( 10 :: nat)) +( 48 :: nat)) # acc1))" (*val stringFromNat : nat -> string*) @@ -58,7 +58,7 @@ fun stringFromNaturalHelper :: " nat \<Rightarrow>(char)list \<Rightarrow>(cha if n =( 0 :: nat) then acc1 else - stringFromNaturalHelper (n div( 10 :: nat)) (char_of_nat ( ((n mod( 10 :: nat)) +( 48 :: nat))) # acc1))" + stringFromNaturalHelper (n div( 10 :: nat)) ((%n. char_of (n::nat)) ( ((n mod( 10 :: nat)) +( 48 :: nat))) # acc1))" (*val stringFromNatural : natural -> string*) @@ -98,7 +98,7 @@ definition nth :: " string \<Rightarrow> nat \<Rightarrow> char " where (*val stringConcat : list string -> string*) definition stringConcat :: "(string)list \<Rightarrow> string " where " stringConcat s = ( - List.foldr (op@) s (''''))" + List.foldr (@) s (''''))" (******************************************************************************) diff --git a/snapshots/isabelle/lib/lem/Lem_tuple.thy b/snapshots/isabelle/lib/lem/Lem_tuple.thy index 66f1a209..292dc17a 100644 --- a/snapshots/isabelle/lib/lem/Lem_tuple.thy +++ b/snapshots/isabelle/lib/lem/Lem_tuple.thy @@ -1,11 +1,11 @@ -chapter \<open>Generated by Lem from tuple.lem.\<close> +chapter \<open>Generated by Lem from \<open>tuple.lem\<close>.\<close> theory "Lem_tuple" -imports - Main - "Lem_bool" - "Lem_basic_classes" +imports + Main + "Lem_bool" + "Lem_basic_classes" begin diff --git a/snapshots/isabelle/lib/lem/Lem_word.thy b/snapshots/isabelle/lib/lem/Lem_word.thy index bc56da3c..af0ef53a 100644 --- a/snapshots/isabelle/lib/lem/Lem_word.thy +++ b/snapshots/isabelle/lib/lem/Lem_word.thy @@ -1,15 +1,15 @@ -chapter \<open>Generated by Lem from word.lem.\<close> +chapter \<open>Generated by Lem from \<open>word.lem\<close>.\<close> theory "Lem_word" -imports - Main - "Lem_bool" - "Lem_maybe" - "Lem_num" - "Lem_basic_classes" - "Lem_list" - "~~/src/HOL/Word/Word" +imports + Main + "Lem_bool" + "Lem_maybe" + "Lem_num" + "Lem_basic_classes" + "Lem_list" + "HOL-Word.Word" begin @@ -17,7 +17,7 @@ begin (*open import Bool Maybe Num Basic_classes List*) -(*open import {isabelle} `~~/src/HOL/Word/Word`*) +(*open import {isabelle} `HOL-Word.Word`*) (*open import {hol} `wordsTheory` `wordsLib`*) @@ -62,8 +62,8 @@ definition bitSeqFromBoolList :: "(bool)list \<Rightarrow>(bitSequence)option " (*val cleanBitSeq : bitSequence -> bitSequence*) fun cleanBitSeq :: " bitSequence \<Rightarrow> bitSequence " where " cleanBitSeq (BitSeq len s bl) = ( (case len of - None => (BitSeq len s (List.rev (dropWhile ((op \<longleftrightarrow>) s) (List.rev bl)))) - | Some n => (BitSeq len s (List.rev (dropWhile ((op \<longleftrightarrow>) s) (List.rev (List.take (n-( 1 :: nat)) bl))))) + None => (BitSeq len s (List.rev (dropWhile ((\<longleftrightarrow>) s) (List.rev bl)))) + | Some n => (BitSeq len s (List.rev (dropWhile ((\<longleftrightarrow>) s) (List.rev (List.take (n-( 1 :: nat)) bl))))) ))" @@ -143,10 +143,10 @@ definition bitSeqBinop :: "(bool \<Rightarrow> bool \<Rightarrow> bool)\<Righta definition bitSeqAnd :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bitSequence " where - " bitSeqAnd = ( bitSeqBinop (op \<and>))" + " bitSeqAnd = ( bitSeqBinop (\<and>))" definition bitSeqOr :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bitSequence " where - " bitSeqOr = ( bitSeqBinop (op \<or>))" + " bitSeqOr = ( bitSeqBinop (\<or>))" definition bitSeqXor :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bitSequence " where " bitSeqXor = ( bitSeqBinop (\<lambda> b1 b2. \<not> (b1 \<longleftrightarrow> b2)))" @@ -277,27 +277,27 @@ definition bitSeqArithBinTest :: "(int \<Rightarrow> int \<Rightarrow> 'a)\<Rig (*val bitSeqLess : bitSequence -> bitSequence -> bool*) definition bitSeqLess :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bool " where - " bitSeqLess bs1 bs2 = ( bitSeqArithBinTest (op<) bs1 bs2 )" + " bitSeqLess bs1 bs2 = ( bitSeqArithBinTest (<) bs1 bs2 )" (*val bitSeqLessEqual : bitSequence -> bitSequence -> bool*) definition bitSeqLessEqual :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bool " where - " bitSeqLessEqual bs1 bs2 = ( bitSeqArithBinTest (op \<le>) bs1 bs2 )" + " bitSeqLessEqual bs1 bs2 = ( bitSeqArithBinTest (\<le>) bs1 bs2 )" (*val bitSeqGreater : bitSequence -> bitSequence -> bool*) definition bitSeqGreater :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bool " where - " bitSeqGreater bs1 bs2 = ( bitSeqArithBinTest (op>) bs1 bs2 )" + " bitSeqGreater bs1 bs2 = ( bitSeqArithBinTest (>) bs1 bs2 )" (*val bitSeqGreaterEqual : bitSequence -> bitSequence -> bool*) definition bitSeqGreaterEqual :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bool " where - " bitSeqGreaterEqual bs1 bs2 = ( bitSeqArithBinTest (op \<ge>) bs1 bs2 )" + " bitSeqGreaterEqual bs1 bs2 = ( bitSeqArithBinTest (\<ge>) bs1 bs2 )" (*val bitSeqCompare : bitSequence -> bitSequence -> ordering*) definition bitSeqCompare :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> ordering " where - " bitSeqCompare bs1 bs2 = ( bitSeqArithBinTest (genericCompare (op<) (op=)) bs1 bs2 )" + " bitSeqCompare bs1 bs2 = ( bitSeqArithBinTest (genericCompare (<) (=)) bs1 bs2 )" definition instance_Basic_classes_Ord_Word_bitSequence_dict :: "(bitSequence)Ord_class " where @@ -329,7 +329,7 @@ definition instance_Num_NumNegate_Word_bitSequence_dict :: "(bitSequence)NumNeg (*val bitSeqAdd : bitSequence -> bitSequence -> bitSequence*) definition bitSeqAdd :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bitSequence " where - " bitSeqAdd bs1 bs2 = ( bitSeqArithBinOp (op+) bs1 bs2 )" + " bitSeqAdd bs1 bs2 = ( bitSeqArithBinOp (+) bs1 bs2 )" definition instance_Num_NumAdd_Word_bitSequence_dict :: "(bitSequence)NumAdd_class " where @@ -340,7 +340,7 @@ definition instance_Num_NumAdd_Word_bitSequence_dict :: "(bitSequence)NumAdd_cl (*val bitSeqMinus : bitSequence -> bitSequence -> bitSequence*) definition bitSeqMinus :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bitSequence " where - " bitSeqMinus bs1 bs2 = ( bitSeqArithBinOp (op-) bs1 bs2 )" + " bitSeqMinus bs1 bs2 = ( bitSeqArithBinOp (-) bs1 bs2 )" definition instance_Num_NumMinus_Word_bitSequence_dict :: "(bitSequence)NumMinus_class " where @@ -373,7 +373,7 @@ definition instance_Num_NumPred_Word_bitSequence_dict :: "(bitSequence)NumPred_ (*val bitSeqMult : bitSequence -> bitSequence -> bitSequence*) definition bitSeqMult :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bitSequence " where - " bitSeqMult bs1 bs2 = ( bitSeqArithBinOp (op*) bs1 bs2 )" + " bitSeqMult bs1 bs2 = ( bitSeqArithBinOp ( * ) bs1 bs2 )" definition instance_Num_NumMult_Word_bitSequence_dict :: "(bitSequence)NumMult_class " where @@ -396,7 +396,7 @@ definition instance_Num_NumPow_Word_bitSequence_dict :: "(bitSequence)NumPow_cl (*val bitSeqDiv : bitSequence -> bitSequence -> bitSequence*) definition bitSeqDiv :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bitSequence " where - " bitSeqDiv bs1 bs2 = ( bitSeqArithBinOp (op div) bs1 bs2 )" + " bitSeqDiv bs1 bs2 = ( bitSeqArithBinOp (div) bs1 bs2 )" definition instance_Num_NumIntegerDivision_Word_bitSequence_dict :: "(bitSequence)NumIntegerDivision_class " where @@ -413,7 +413,7 @@ definition instance_Num_NumDivision_Word_bitSequence_dict :: "(bitSequence)NumD (*val bitSeqMod : bitSequence -> bitSequence -> bitSequence*) definition bitSeqMod :: " bitSequence \<Rightarrow> bitSequence \<Rightarrow> bitSequence " where - " bitSeqMod bs1 bs2 = ( bitSeqArithBinOp (op mod) bs1 bs2 )" + " bitSeqMod bs1 bs2 = ( bitSeqArithBinOp (mod) bs1 bs2 )" definition instance_Num_NumRemainder_Word_bitSequence_dict :: "(bitSequence)NumRemainder_class " where @@ -555,7 +555,7 @@ definition instance_Word_WordNot_Num_int32_dict :: "( 32 word)WordNot_class " definition instance_Word_WordOr_Num_int32_dict :: "( 32 word)WordOr_class " where " instance_Word_WordOr_Num_int32_dict = ((| - lor_method = (op OR)|) )" + lor_method = (OR)|) )" (*val int32Lxor : int32 -> int32 -> int32*) (* XXX: fix *) @@ -563,7 +563,7 @@ definition instance_Word_WordOr_Num_int32_dict :: "( 32 word)WordOr_class " wh definition instance_Word_WordXor_Num_int32_dict :: "( 32 word)WordXor_class " where " instance_Word_WordXor_Num_int32_dict = ((| - lxor_method = (op XOR)|) )" + lxor_method = (XOR)|) )" (*val int32Land : int32 -> int32 -> int32*) (* XXX: fix *) @@ -571,7 +571,7 @@ definition instance_Word_WordXor_Num_int32_dict :: "( 32 word)WordXor_class " definition instance_Word_WordAnd_Num_int32_dict :: "( 32 word)WordAnd_class " where " instance_Word_WordAnd_Num_int32_dict = ((| - land_method = (op AND)|) )" + land_method = (AND)|) )" (*val int32Lsl : int32 -> nat -> int32*) (* XXX: fix *) @@ -579,7 +579,7 @@ definition instance_Word_WordAnd_Num_int32_dict :: "( 32 word)WordAnd_class " definition instance_Word_WordLsl_Num_int32_dict :: "( 32 word)WordLsl_class " where " instance_Word_WordLsl_Num_int32_dict = ((| - lsl_method = (op<<)|) )" + lsl_method = (<<)|) )" (*val int32Lsr : int32 -> nat -> int32*) (* XXX: fix *) @@ -587,7 +587,7 @@ definition instance_Word_WordLsl_Num_int32_dict :: "( 32 word)WordLsl_class " definition instance_Word_WordLsr_Num_int32_dict :: "( 32 word)WordLsr_class " where " instance_Word_WordLsr_Num_int32_dict = ((| - lsr_method = (op>>)|) )" + lsr_method = (>>)|) )" @@ -596,7 +596,7 @@ definition instance_Word_WordLsr_Num_int32_dict :: "( 32 word)WordLsr_class " definition instance_Word_WordAsr_Num_int32_dict :: "( 32 word)WordAsr_class " where " instance_Word_WordAsr_Num_int32_dict = ((| - asr_method = (op>>>)|) )" + asr_method = (>>>)|) )" @@ -617,7 +617,7 @@ definition instance_Word_WordNot_Num_int64_dict :: "( 64 word)WordNot_class " definition instance_Word_WordOr_Num_int64_dict :: "( 64 word)WordOr_class " where " instance_Word_WordOr_Num_int64_dict = ((| - lor_method = (op OR)|) )" + lor_method = (OR)|) )" (*val int64Lxor : int64 -> int64 -> int64*) (* XXX: fix *) @@ -625,7 +625,7 @@ definition instance_Word_WordOr_Num_int64_dict :: "( 64 word)WordOr_class " wh definition instance_Word_WordXor_Num_int64_dict :: "( 64 word)WordXor_class " where " instance_Word_WordXor_Num_int64_dict = ((| - lxor_method = (op XOR)|) )" + lxor_method = (XOR)|) )" (*val int64Land : int64 -> int64 -> int64*) (* XXX: fix *) @@ -633,7 +633,7 @@ definition instance_Word_WordXor_Num_int64_dict :: "( 64 word)WordXor_class " definition instance_Word_WordAnd_Num_int64_dict :: "( 64 word)WordAnd_class " where " instance_Word_WordAnd_Num_int64_dict = ((| - land_method = (op AND)|) )" + land_method = (AND)|) )" (*val int64Lsl : int64 -> nat -> int64*) (* XXX: fix *) @@ -641,7 +641,7 @@ definition instance_Word_WordAnd_Num_int64_dict :: "( 64 word)WordAnd_class " definition instance_Word_WordLsl_Num_int64_dict :: "( 64 word)WordLsl_class " where " instance_Word_WordLsl_Num_int64_dict = ((| - lsl_method = (op<<)|) )" + lsl_method = (<<)|) )" (*val int64Lsr : int64 -> nat -> int64*) (* XXX: fix *) @@ -649,7 +649,7 @@ definition instance_Word_WordLsl_Num_int64_dict :: "( 64 word)WordLsl_class " definition instance_Word_WordLsr_Num_int64_dict :: "( 64 word)WordLsr_class " where " instance_Word_WordLsr_Num_int64_dict = ((| - lsr_method = (op>>)|) )" + lsr_method = (>>)|) )" (*val int64Asr : int64 -> nat -> int64*) (* XXX: fix *) @@ -657,7 +657,7 @@ definition instance_Word_WordLsr_Num_int64_dict :: "( 64 word)WordLsr_class " definition instance_Word_WordAsr_Num_int64_dict :: "( 64 word)WordAsr_class " where " instance_Word_WordAsr_Num_int64_dict = ((| - asr_method = (op>>>)|) )" + asr_method = (>>>)|) )" diff --git a/snapshots/isabelle/lib/sail/Sail2_instr_kinds.thy b/snapshots/isabelle/lib/sail/Sail2_instr_kinds.thy index d8f3b593..b9381c39 100644 --- a/snapshots/isabelle/lib/sail/Sail2_instr_kinds.thy +++ b/snapshots/isabelle/lib/sail/Sail2_instr_kinds.thy @@ -1,10 +1,10 @@ -chapter \<open>Generated by Lem from ../../src/lem_interp/sail2_instr_kinds.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/lem_interp/sail2_instr_kinds.lem\<close>.\<close> theory "Sail2_instr_kinds" -imports - Main - "Lem_pervasives_extra" +imports + Main + "LEM.Lem_pervasives_extra" begin @@ -21,7 +21,7 @@ record 'a EnumerationType_class= (*val enumeration_typeCompare : forall 'a. EnumerationType 'a => 'a -> 'a -> ordering*) definition enumeration_typeCompare :: " 'a EnumerationType_class \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ordering " where " enumeration_typeCompare dict_Sail2_instr_kinds_EnumerationType_a e1 e2 = ( - (genericCompare (op<) (op=) ( + (genericCompare (<) (=) ( (toNat_method dict_Sail2_instr_kinds_EnumerationType_a) e1) ((toNat_method dict_Sail2_instr_kinds_EnumerationType_a) e2)))" diff --git a/snapshots/isabelle/lib/sail/Sail2_operators.thy b/snapshots/isabelle/lib/sail/Sail2_operators.thy index fc4d4749..9cd11709 100644 --- a/snapshots/isabelle/lib/sail/Sail2_operators.thy +++ b/snapshots/isabelle/lib/sail/Sail2_operators.thy @@ -1,12 +1,12 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_operators.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_operators.lem\<close>.\<close> theory "Sail2_operators" -imports - Main - "Lem_pervasives_extra" - "Lem_machine_word" - "Sail2_values" +imports + Main + "LEM.Lem_pervasives_extra" + "LEM.Lem_machine_word" + "Sail2_values" begin @@ -253,10 +253,10 @@ definition quots_bv :: " 'a Bitvector_class \<Rightarrow> 'b Bitvector_class \< definition mod_mword :: "('a::len)Word.word \<Rightarrow>('a::len)Word.word \<Rightarrow>('a::len)Word.word " where - " mod_mword = ( (op mod))" + " mod_mword = ( (mod))" definition quot_mword :: "('a::len)Word.word \<Rightarrow>('a::len)Word.word \<Rightarrow>('a::len)Word.word " where - " quot_mword = ( (op div))" + " quot_mword = ( (div))" definition quots_mword :: "('a::len)Word.word \<Rightarrow>('a::len)Word.word \<Rightarrow>('a::len)Word.word " where " quots_mword = ( Lem_machine_word.signedDivide )" diff --git a/snapshots/isabelle/lib/sail/Sail2_operators_bitlists.thy b/snapshots/isabelle/lib/sail/Sail2_operators_bitlists.thy index 35f43769..6ce1a5cf 100644 --- a/snapshots/isabelle/lib/sail/Sail2_operators_bitlists.thy +++ b/snapshots/isabelle/lib/sail/Sail2_operators_bitlists.thy @@ -1,15 +1,15 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_operators_bitlists.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_operators_bitlists.lem\<close>.\<close> theory "Sail2_operators_bitlists" -imports - Main - "Lem_pervasives_extra" - "Lem_machine_word" - "Sail2_values" - "Sail2_operators" - "Sail2_prompt_monad" - "Sail2_prompt" +imports + Main + "LEM.Lem_pervasives_extra" + "LEM.Lem_machine_word" + "Sail2_values" + "Sail2_operators" + "Sail2_prompt_monad" + "Sail2_prompt" begin @@ -311,26 +311,26 @@ definition arith_op_double_bl :: " 'a Bitvector_class \<Rightarrow>(int \<Right (*val mult_vec : list bitU -> list bitU -> list bitU*) (*val mults_vec : list bitU -> list bitU -> list bitU*) definition add_vec :: "(bitU)list \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where - " add_vec = ( (\<lambda> l r. List.map (\<lambda> b. b) (arith_op_bits (op+) False (List.map (\<lambda> b. b) l) (List.map (\<lambda> b. b) r))))" + " add_vec = ( (\<lambda> l r. List.map (\<lambda> b. b) (arith_op_bits (+) False (List.map (\<lambda> b. b) l) (List.map (\<lambda> b. b) r))))" definition adds_vec :: "(bitU)list \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where - " adds_vec = ( (\<lambda> l r. List.map (\<lambda> b. b) (arith_op_bits (op+) True (List.map (\<lambda> b. b) l) (List.map (\<lambda> b. b) r))))" + " adds_vec = ( (\<lambda> l r. List.map (\<lambda> b. b) (arith_op_bits (+) True (List.map (\<lambda> b. b) l) (List.map (\<lambda> b. b) r))))" definition sub_vec :: "(bitU)list \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where - " sub_vec = ( (\<lambda> l r. List.map (\<lambda> b. b) (arith_op_bits (op-) False (List.map (\<lambda> b. b) l) (List.map (\<lambda> b. b) r))))" + " sub_vec = ( (\<lambda> l r. List.map (\<lambda> b. b) (arith_op_bits (-) False (List.map (\<lambda> b. b) l) (List.map (\<lambda> b. b) r))))" definition subs_vec :: "(bitU)list \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where - " subs_vec = ( (\<lambda> l r. List.map (\<lambda> b. b) (arith_op_bits (op-) True (List.map (\<lambda> b. b) l) (List.map (\<lambda> b. b) r))))" + " subs_vec = ( (\<lambda> l r. List.map (\<lambda> b. b) (arith_op_bits (-) True (List.map (\<lambda> b. b) l) (List.map (\<lambda> b. b) r))))" definition mult_vec :: "(bitU)list \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where " mult_vec = ( arith_op_double_bl (instance_Sail2_values_Bitvector_list_dict - instance_Sail2_values_BitU_Sail2_values_bitU_dict) (op*) False )" + instance_Sail2_values_BitU_Sail2_values_bitU_dict) ( * ) False )" definition mults_vec :: "(bitU)list \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where " mults_vec = ( arith_op_double_bl (instance_Sail2_values_Bitvector_list_dict - instance_Sail2_values_BitU_Sail2_values_bitU_dict) (op*) True )" + instance_Sail2_values_BitU_Sail2_values_bitU_dict) ( * ) True )" (*val add_vec_int : list bitU -> integer -> list bitU*) @@ -339,17 +339,17 @@ definition mults_vec :: "(bitU)list \<Rightarrow>(bitU)list \<Rightarrow>(bitU) definition add_vec_int :: "(bitU)list \<Rightarrow> int \<Rightarrow>(bitU)list " where " add_vec_int l r = ( arith_op_bv_int (instance_Sail2_values_Bitvector_list_dict - instance_Sail2_values_BitU_Sail2_values_bitU_dict) (op+) False l r )" + instance_Sail2_values_BitU_Sail2_values_bitU_dict) (+) False l r )" definition sub_vec_int :: "(bitU)list \<Rightarrow> int \<Rightarrow>(bitU)list " where " sub_vec_int l r = ( arith_op_bv_int (instance_Sail2_values_Bitvector_list_dict - instance_Sail2_values_BitU_Sail2_values_bitU_dict) (op-) False l r )" + instance_Sail2_values_BitU_Sail2_values_bitU_dict) (-) False l r )" definition mult_vec_int :: "(bitU)list \<Rightarrow> int \<Rightarrow>(bitU)list " where " mult_vec_int l r = ( arith_op_double_bl (instance_Sail2_values_Bitvector_list_dict - instance_Sail2_values_BitU_Sail2_values_bitU_dict) (op*) False l (List.map (\<lambda> b. b) (bits_of_int (int (List.length l)) r)))" + instance_Sail2_values_BitU_Sail2_values_bitU_dict) ( * ) False l (List.map (\<lambda> b. b) (bits_of_int (int (List.length l)) r)))" (*val add_int_vec : integer -> list bitU -> list bitU*) @@ -358,17 +358,17 @@ definition mult_vec_int :: "(bitU)list \<Rightarrow> int \<Rightarrow>(bitU)lis definition add_int_vec :: " int \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where " add_int_vec l r = ( arith_op_int_bv (instance_Sail2_values_Bitvector_list_dict - instance_Sail2_values_BitU_Sail2_values_bitU_dict) (op+) False l r )" + instance_Sail2_values_BitU_Sail2_values_bitU_dict) (+) False l r )" definition sub_int_vec :: " int \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where " sub_int_vec l r = ( arith_op_int_bv (instance_Sail2_values_Bitvector_list_dict - instance_Sail2_values_BitU_Sail2_values_bitU_dict) (op-) False l r )" + instance_Sail2_values_BitU_Sail2_values_bitU_dict) (-) False l r )" definition mult_int_vec :: " int \<Rightarrow>(bitU)list \<Rightarrow>(bitU)list " where " mult_int_vec l r = ( arith_op_double_bl (instance_Sail2_values_Bitvector_list_dict - instance_Sail2_values_BitU_Sail2_values_bitU_dict) (op*) False (List.map (\<lambda> b. b) (bits_of_int (int (List.length r)) l)) r )" + instance_Sail2_values_BitU_Sail2_values_bitU_dict) ( * ) False (List.map (\<lambda> b. b) (bits_of_int (int (List.length r)) l)) r )" (*val add_vec_bit : list bitU -> bitU -> list bitU*) @@ -378,11 +378,11 @@ definition mult_int_vec :: " int \<Rightarrow>(bitU)list \<Rightarrow>(bitU)lis definition add_vec_bool :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> 'a " where " add_vec_bool dict_Sail2_values_Bitvector_a l r = ( arith_op_bv_bool - dict_Sail2_values_Bitvector_a (op+) False l r )" + dict_Sail2_values_Bitvector_a (+) False l r )" definition add_vec_bit_maybe :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bitU \<Rightarrow> 'a option " where " add_vec_bit_maybe dict_Sail2_values_Bitvector_a l r = ( arith_op_bv_bit - dict_Sail2_values_Bitvector_a (op+) False l r )" + dict_Sail2_values_Bitvector_a (+) False l r )" definition add_vec_bit_fail :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bitU \<Rightarrow>('d,'a,'c)monad " where " add_vec_bit_fail dict_Sail2_values_Bitvector_a l r = ( maybe_fail (''add_vec_bit'') (add_vec_bit_maybe @@ -400,11 +400,11 @@ definition add_vec_bit :: "(bitU)list \<Rightarrow> bitU \<Rightarrow>(bitU)lis definition adds_vec_bool :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> 'a " where " adds_vec_bool dict_Sail2_values_Bitvector_a l r = ( arith_op_bv_bool - dict_Sail2_values_Bitvector_a (op+) True l r )" + dict_Sail2_values_Bitvector_a (+) True l r )" definition adds_vec_bit_maybe :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bitU \<Rightarrow> 'a option " where " adds_vec_bit_maybe dict_Sail2_values_Bitvector_a l r = ( arith_op_bv_bit - dict_Sail2_values_Bitvector_a (op+) True l r )" + dict_Sail2_values_Bitvector_a (+) True l r )" definition adds_vec_bit_fail :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bitU \<Rightarrow>('d,'a,'c)monad " where " adds_vec_bit_fail dict_Sail2_values_Bitvector_a l r = ( maybe_fail (''adds_vec_bit'') (adds_vec_bit_maybe @@ -422,11 +422,11 @@ definition adds_vec_bit :: "(bitU)list \<Rightarrow> bitU \<Rightarrow>(bitU)li definition sub_vec_bool :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> 'a " where " sub_vec_bool dict_Sail2_values_Bitvector_a l r = ( arith_op_bv_bool - dict_Sail2_values_Bitvector_a (op-) False l r )" + dict_Sail2_values_Bitvector_a (-) False l r )" definition sub_vec_bit_maybe :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bitU \<Rightarrow> 'a option " where " sub_vec_bit_maybe dict_Sail2_values_Bitvector_a l r = ( arith_op_bv_bit - dict_Sail2_values_Bitvector_a (op-) False l r )" + dict_Sail2_values_Bitvector_a (-) False l r )" definition sub_vec_bit_fail :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bitU \<Rightarrow>('d,'a,'c)monad " where " sub_vec_bit_fail dict_Sail2_values_Bitvector_a l r = ( maybe_fail (''sub_vec_bit'') (sub_vec_bit_maybe @@ -444,11 +444,11 @@ definition sub_vec_bit :: "(bitU)list \<Rightarrow> bitU \<Rightarrow>(bitU)lis definition subs_vec_bool :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> 'a " where " subs_vec_bool dict_Sail2_values_Bitvector_a l r = ( arith_op_bv_bool - dict_Sail2_values_Bitvector_a (op-) True l r )" + dict_Sail2_values_Bitvector_a (-) True l r )" definition subs_vec_bit_maybe :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bitU \<Rightarrow> 'a option " where " subs_vec_bit_maybe dict_Sail2_values_Bitvector_a l r = ( arith_op_bv_bit - dict_Sail2_values_Bitvector_a (op-) True l r )" + dict_Sail2_values_Bitvector_a (-) True l r )" definition subs_vec_bit_fail :: " 'a Bitvector_class \<Rightarrow> 'a \<Rightarrow> bitU \<Rightarrow>('d,'a,'c)monad " where " subs_vec_bit_fail dict_Sail2_values_Bitvector_a l r = ( maybe_fail (''sub_vec_bit'') (subs_vec_bit_maybe diff --git a/snapshots/isabelle/lib/sail/Sail2_operators_mwords.thy b/snapshots/isabelle/lib/sail/Sail2_operators_mwords.thy index 44d29a7f..8650d968 100644 --- a/snapshots/isabelle/lib/sail/Sail2_operators_mwords.thy +++ b/snapshots/isabelle/lib/sail/Sail2_operators_mwords.thy @@ -1,15 +1,15 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_operators_mwords.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_operators_mwords.lem\<close>.\<close> theory "Sail2_operators_mwords" -imports - Main - "Lem_pervasives_extra" - "Lem_machine_word" - "Sail2_values" - "Sail2_operators" - "Sail2_prompt_monad" - "Sail2_prompt" +imports + Main + "LEM.Lem_pervasives_extra" + "LEM.Lem_machine_word" + "Sail2_values" + "Sail2_operators" + "Sail2_prompt_monad" + "Sail2_prompt" begin @@ -288,15 +288,15 @@ definition mults_vec :: "('a::len)Word.word \<Rightarrow>('a::len)Word.word \<R (*val mult_vec_int : forall 'a 'b. Size 'a, Size 'b => mword 'a -> integer -> mword 'b*) definition add_vec_int :: "('a::len)Word.word \<Rightarrow> int \<Rightarrow>('a::len)Word.word " where " add_vec_int l r = ( arith_op_bv_int - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op+) False l r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict (+) False l r )" definition sub_vec_int :: "('a::len)Word.word \<Rightarrow> int \<Rightarrow>('a::len)Word.word " where " sub_vec_int l r = ( arith_op_bv_int - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op-) False l r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict (-) False l r )" definition mult_vec_int :: "('a::len)Word.word \<Rightarrow> int \<Rightarrow>('b::len)Word.word " where " mult_vec_int l r = ( arith_op_bv_int - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op*) False (Word.ucast l :: ( 'b::len)Word.word) r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict ( * ) False (Word.ucast l :: ( 'b::len)Word.word) r )" (*val add_int_vec : forall 'a. Size 'a => integer -> mword 'a -> mword 'a*) @@ -304,15 +304,15 @@ definition mult_vec_int :: "('a::len)Word.word \<Rightarrow> int \<Rightarrow>( (*val mult_int_vec : forall 'a 'b. Size 'a, Size 'b => integer -> mword 'a -> mword 'b*) definition add_int_vec :: " int \<Rightarrow>('a::len)Word.word \<Rightarrow>('a::len)Word.word " where " add_int_vec l r = ( arith_op_int_bv - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op+) False l r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict (+) False l r )" definition sub_int_vec :: " int \<Rightarrow>('a::len)Word.word \<Rightarrow>('a::len)Word.word " where " sub_int_vec l r = ( arith_op_int_bv - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op-) False l r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict (-) False l r )" definition mult_int_vec :: " int \<Rightarrow>('a::len)Word.word \<Rightarrow>('b::len)Word.word " where " mult_int_vec l r = ( arith_op_int_bv - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op*) False l (Word.ucast r :: ( 'b::len)Word.word))" + instance_Sail2_values_Bitvector_Machine_word_mword_dict ( * ) False l (Word.ucast r :: ( 'b::len)Word.word))" (*val add_vec_bool : forall 'a. Size 'a => mword 'a -> bool -> mword 'a*) @@ -322,7 +322,7 @@ definition mult_int_vec :: " int \<Rightarrow>('a::len)Word.word \<Rightarrow>( definition add_vec_bool :: "('a::len)Word.word \<Rightarrow> bool \<Rightarrow>('a::len)Word.word " where " add_vec_bool l r = ( arith_op_bv_bool - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op+) False l r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict (+) False l r )" definition add_vec_bit_maybe :: "('a::len)Word.word \<Rightarrow> bitU \<Rightarrow>(('a::len)Word.word)option " where " add_vec_bit_maybe l r = ( map_option (add_vec_bool l) (bool_of_bitU r))" @@ -339,7 +339,7 @@ definition add_vec_bit :: "('a::len)Word.word \<Rightarrow> bitU \<Rightarrow>( definition adds_vec_bool :: "('a::len)Word.word \<Rightarrow> bool \<Rightarrow>('a::len)Word.word " where " adds_vec_bool l r = ( arith_op_bv_bool - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op+) True l r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict (+) True l r )" definition adds_vec_bit_maybe :: "('a::len)Word.word \<Rightarrow> bitU \<Rightarrow>(('a::len)Word.word)option " where " adds_vec_bit_maybe l r = ( map_option (adds_vec_bool l) (bool_of_bitU r))" @@ -356,7 +356,7 @@ definition adds_vec_bit :: "('a::len)Word.word \<Rightarrow> bitU \<Rightarrow> definition sub_vec_bool :: "('a::len)Word.word \<Rightarrow> bool \<Rightarrow>('a::len)Word.word " where " sub_vec_bool l r = ( arith_op_bv_bool - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op-) False l r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict (-) False l r )" definition sub_vec_bit_maybe :: "('a::len)Word.word \<Rightarrow> bitU \<Rightarrow>(('a::len)Word.word)option " where " sub_vec_bit_maybe l r = ( map_option (sub_vec_bool l) (bool_of_bitU r))" @@ -373,7 +373,7 @@ definition sub_vec_bit :: "('a::len)Word.word \<Rightarrow> bitU \<Rightarrow>( definition subs_vec_bool :: "('a::len)Word.word \<Rightarrow> bool \<Rightarrow>('a::len)Word.word " where " subs_vec_bool l r = ( arith_op_bv_bool - instance_Sail2_values_Bitvector_Machine_word_mword_dict (op-) True l r )" + instance_Sail2_values_Bitvector_Machine_word_mword_dict (-) True l r )" definition subs_vec_bit_maybe :: "('a::len)Word.word \<Rightarrow> bitU \<Rightarrow>(('a::len)Word.word)option " where " subs_vec_bit_maybe l r = ( map_option (subs_vec_bool l) (bool_of_bitU r))" diff --git a/snapshots/isabelle/lib/sail/Sail2_prompt.thy b/snapshots/isabelle/lib/sail/Sail2_prompt.thy index e639d59a..4f9b090b 100644 --- a/snapshots/isabelle/lib/sail/Sail2_prompt.thy +++ b/snapshots/isabelle/lib/sail/Sail2_prompt.thy @@ -1,13 +1,13 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_prompt.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_prompt.lem\<close>.\<close> theory "Sail2_prompt" -imports - Main - "Lem_pervasives_extra" - "Sail2_values" - "Sail2_prompt_monad" - "Sail2_prompt_monad_lemmas" +imports + Main + "LEM.Lem_pervasives_extra" + "Sail2_values" + "Sail2_prompt_monad" + "Sail2_prompt_monad_lemmas" begin diff --git a/snapshots/isabelle/lib/sail/Sail2_prompt_monad.thy b/snapshots/isabelle/lib/sail/Sail2_prompt_monad.thy index d1259332..212f9b9c 100644 --- a/snapshots/isabelle/lib/sail/Sail2_prompt_monad.thy +++ b/snapshots/isabelle/lib/sail/Sail2_prompt_monad.thy @@ -1,12 +1,12 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_prompt_monad.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_prompt_monad.lem\<close>.\<close> theory "Sail2_prompt_monad" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" +imports + Main + "LEM.Lem_pervasives_extra" + "Sail2_instr_kinds" + "Sail2_values" begin @@ -87,7 +87,7 @@ definition undefined_bool :: " unit \<Rightarrow>('rv,(bool),'e)monad " where (*val assert_exp : forall 'rv 'e. bool -> string -> monad 'rv unit 'e*) definition assert_exp :: " bool \<Rightarrow> string \<Rightarrow>('rv,(unit),'e)monad " where - " assert_exp exp msg = ( if exp then Done () else Fail msg )" + " assert_exp exp1 msg = ( if exp1 then Done () else Fail msg )" (*val throw : forall 'rv 'a 'e. 'e -> monad 'rv 'a 'e*) diff --git a/snapshots/isabelle/lib/sail/Sail2_state.thy b/snapshots/isabelle/lib/sail/Sail2_state.thy index dd591a86..82c86c64 100644 --- a/snapshots/isabelle/lib/sail/Sail2_state.thy +++ b/snapshots/isabelle/lib/sail/Sail2_state.thy @@ -1,13 +1,13 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_state.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_state.lem\<close>.\<close> theory "Sail2_state" -imports - Main - "Lem_pervasives_extra" - "Sail2_values" - "Sail2_state_monad" - "Sail2_state_monad_lemmas" +imports + Main + "LEM.Lem_pervasives_extra" + "Sail2_values" + "Sail2_state_monad" + "Sail2_state_monad_lemmas" begin diff --git a/snapshots/isabelle/lib/sail/Sail2_state_lifting.thy b/snapshots/isabelle/lib/sail/Sail2_state_lifting.thy index 957a7940..c10d4651 100644 --- a/snapshots/isabelle/lib/sail/Sail2_state_lifting.thy +++ b/snapshots/isabelle/lib/sail/Sail2_state_lifting.thy @@ -1,15 +1,15 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_state_lifting.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_state_lifting.lem\<close>.\<close> theory "Sail2_state_lifting" -imports - Main - "Lem_pervasives_extra" - "Sail2_values" - "Sail2_prompt_monad" - "Sail2_prompt" - "Sail2_state_monad" - "Sail2_state_monad_lemmas" +imports + Main + "LEM.Lem_pervasives_extra" + "Sail2_values" + "Sail2_prompt_monad" + "Sail2_prompt" + "Sail2_state_monad" + "Sail2_state_monad_lemmas" begin diff --git a/snapshots/isabelle/lib/sail/Sail2_state_monad.thy b/snapshots/isabelle/lib/sail/Sail2_state_monad.thy index e1bf0048..b5e915ac 100644 --- a/snapshots/isabelle/lib/sail/Sail2_state_monad.thy +++ b/snapshots/isabelle/lib/sail/Sail2_state_monad.thy @@ -1,12 +1,12 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_state_monad.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_state_monad.lem\<close>.\<close> theory "Sail2_state_monad" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" +imports + Main + "LEM.Lem_pervasives_extra" + "Sail2_instr_kinds" + "Sail2_values" begin @@ -121,7 +121,7 @@ definition try_catchS :: "('regs sequential_state \<Rightarrow>(('a,'e1)result* (*val assert_expS : forall 'regs 'e. bool -> string -> monadS 'regs unit 'e*) definition assert_expS :: " bool \<Rightarrow> string \<Rightarrow> 'regs sequential_state \<Rightarrow>(((unit),'e)result*'regs sequential_state)set " where - " assert_expS exp msg = ( if exp then returnS () else failS msg )" + " assert_expS exp1 msg = ( if exp1 then returnS () else failS msg )" (* For early return, we abuse exceptions by throwing and catching diff --git a/snapshots/isabelle/lib/sail/Sail2_string.thy b/snapshots/isabelle/lib/sail/Sail2_string.thy index fd2c1d04..18d9f7ec 100644 --- a/snapshots/isabelle/lib/sail/Sail2_string.thy +++ b/snapshots/isabelle/lib/sail/Sail2_string.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_string.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_string.lem\<close>.\<close> theory "Sail2_string" -imports - Main - "Lem_pervasives" - "Lem_list" - "Lem_list_extra" - "Lem_string" - "Lem_string_extra" - "Sail2_operators" - "Sail2_values" +imports + Main + "LEM.Lem_pervasives" + "LEM.Lem_list" + "LEM.Lem_list_extra" + "LEM.Lem_string" + "LEM.Lem_string_extra" + "Sail2_operators" + "Sail2_values" begin @@ -48,7 +48,7 @@ definition string_length :: " string \<Rightarrow> int " where definition string_append :: " string \<Rightarrow> string \<Rightarrow> string " where - " string_append = ( (op@))" + " string_append = ( (@))" (*********************************************** diff --git a/snapshots/isabelle/lib/sail/Sail2_values.thy b/snapshots/isabelle/lib/sail/Sail2_values.thy index 6278f8e3..3bfeac0c 100644 --- a/snapshots/isabelle/lib/sail/Sail2_values.thy +++ b/snapshots/isabelle/lib/sail/Sail2_values.thy @@ -1,11 +1,11 @@ -chapter \<open>Generated by Lem from ../../src/gen_lib/sail2_values.lem.\<close> +chapter \<open>Generated by Lem from \<open>../../src/gen_lib/sail2_values.lem\<close>.\<close> theory "Sail2_values" -imports - Main - "Lem_pervasives_extra" - "Lem_machine_word" +imports + Main + "LEM.Lem_pervasives_extra" + "LEM.Lem_machine_word" begin diff --git a/snapshots/isabelle/riscv/ROOT b/snapshots/isabelle/riscv/ROOT index cfc7f5bd..ea74bca1 100644 --- a/snapshots/isabelle/riscv/ROOT +++ b/snapshots/isabelle/riscv/ROOT @@ -2,3 +2,8 @@ session "Sail-RISC-V" = "Sail" + options [document = false] theories Riscv_lemmas + +session "Sail-RISC-V-Duopod" = "Sail" + + options [document = false] + theories + Riscv_duopod_lemmas diff --git a/snapshots/isabelle/riscv/Riscv.thy b/snapshots/isabelle/riscv/Riscv.thy index 754caaf5..f047010c 100644 --- a/snapshots/isabelle/riscv/Riscv.thy +++ b/snapshots/isabelle/riscv/Riscv.thy @@ -1,18 +1,18 @@ -chapter \<open>Generated by Lem from riscv.lem.\<close> +chapter \<open>Generated by Lem from \<open>riscv.lem\<close>.\<close> theory "Riscv" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_operators_mwords" - "Sail2_prompt_monad" - "Sail2_prompt" - "Sail2_string" - "Riscv_types" - "Riscv_extras" +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" + "Riscv_extras" begin @@ -81,7 +81,7 @@ definition def_spc_backwards :: " string \<Rightarrow> unit " where (*val __raw_SetSlice_int : forall 'w. integer -> ii -> ii -> bits 'w -> ii*) -(*val __GetSlice_int : forall 'n. Size 'n => integer -> ii -> ii -> mword 'n*) +(*val __GetSlice_int : forall 'n . Size 'n => integer -> ii -> ii -> mword 'n*) definition GetSlice_int :: " int \<Rightarrow> int \<Rightarrow> int \<Rightarrow>('n::len)Word.word " where " GetSlice_int n m o1 = ( (get_slice_int n m o1 :: ( 'n::len)Word.word))" @@ -89,7 +89,7 @@ definition GetSlice_int :: " int \<Rightarrow> int \<Rightarrow> int \<Rightarr (*val __raw_SetSlice_bits : forall 'n 'w. integer -> integer -> bits 'n -> ii -> bits 'w -> bits 'n*) -(*val __raw_GetSlice_bits : forall 'n 'w. integer -> integer -> bits 'n -> ii -> bits 'w*) +(*val __raw_GetSlice_bits : forall 'n 'w . integer -> integer -> bits 'n -> ii -> bits 'w*) (*val cast_unit_vec : bitU -> mword ty1*) @@ -110,13 +110,12 @@ definition RISCV_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_ (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] - :: 64 Word.word) addr data \<then> - return True )" + :: 64 Word.word) addr data )" (*val __TraceMemoryWrite : forall 'int8_times_n 'm. integer -> bits 'm -> bits 'int8_times_n -> unit*) -(*val __RISCV_read : forall 'int8_times_n. Size 'int8_times_n => mword ty64 -> integer -> bool -> bool -> bool -> M (maybe (mword 'int8_times_n))*) +(*val __RISCV_read : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> bool -> bool -> bool -> M (maybe (mword 'int8_times_n))*) fun RISCV_read :: "(64)Word.word \<Rightarrow> int \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)option),(exception))monad " where " RISCV_read addr width False False False = ( @@ -203,9 +202,9 @@ definition EXTZ :: " int \<Rightarrow>('n::len)Word.word \<Rightarrow>('m::len) " EXTZ (m__tv :: int) v = ( (zero_extend v m__tv :: ( 'm::len)Word.word))" -(*val zopz0zI_s : forall 'n. Size 'n => mword 'n -> mword 'n -> bool*) +(*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 zopz0zKzJ_s : forall 'n . Size 'n => mword 'n -> mword 'n -> bool*) (*val zopz0zI_u : forall 'n. Size 'n => mword 'n -> mword 'n -> bool*) @@ -252,7 +251,7 @@ definition vector64 :: " int \<Rightarrow>(64)Word.word " where " vector64 n = ( (get_slice_int (( 64 :: int)::ii) n (( 0 :: int)::ii) :: 64 Word.word))" -(*val to_bits : forall 'l. Size 'l => integer -> ii -> mword 'l*) +(*val to_bits : forall 'l . Size 'l => integer -> ii -> mword 'l*) definition to_bits :: " int \<Rightarrow> int \<Rightarrow>('l::len)Word.word " where " to_bits l n = ( (get_slice_int l n (( 0 :: int)::ii) :: ( 'l::len)Word.word))" @@ -398,9 +397,9 @@ definition wX :: " int \<Rightarrow>(64)Word.word \<Rightarrow>((register_value else write_reg x31_ref v) \<then> return (if (((r \<noteq> (( 0 :: int)::ii)))) then print_endline - (((op@) (''x'') - (((op@) ((stringFromInteger r)) - (((op@) ('' <- '') ((string_of_bits v)))))))) + (((@) (''x'') + (((@) ((stringFromInteger r)) + (((@) ('' <- '') ((string_of_bits v)))))))) else () )))" @@ -607,6 +606,70 @@ fun readType_to_str :: " ReadType \<Rightarrow> string " where |" readType_to_str Data = ( (''D''))" +(*val word_width_of_num : integer -> word_width*) + +definition word_width_of_num :: " int \<Rightarrow> word_width " where + " word_width_of_num arg0 = ( + (let p00 = arg0 in + if (((p00 = (( 0 :: int)::ii)))) then BYTE + else if (((p00 = (( 1 :: int)::ii)))) then HALF + else if (((p00 = (( 2 :: int)::ii)))) then WORD + else DOUBLE))" + + +(*val num_of_word_width : word_width -> integer*) + +fun num_of_word_width :: " word_width \<Rightarrow> 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 \<Rightarrow> InterruptType " where + " InterruptType_of_num arg0 = ( + (let p00 = arg0 in + if (((p00 = (( 0 :: int)::ii)))) then I_U_Software + else if (((p00 = (( 1 :: int)::ii)))) then I_S_Software + else if (((p00 = (( 2 :: int)::ii)))) then I_M_Software + else if (((p00 = (( 3 :: int)::ii)))) then I_U_Timer + else if (((p00 = (( 4 :: int)::ii)))) then I_S_Timer + else if (((p00 = (( 5 :: int)::ii)))) then I_M_Timer + else if (((p00 = (( 6 :: int)::ii)))) then I_U_External + else if (((p00 = (( 7 :: int)::ii)))) then I_S_External + else I_M_External))" + + +(*val num_of_InterruptType : InterruptType -> integer*) + +fun num_of_InterruptType :: " InterruptType \<Rightarrow> 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 ty4*) + +fun interruptType_to_bits :: " InterruptType \<Rightarrow>(4)Word.word " where + " interruptType_to_bits I_U_Software = ( (vec_of_bits [B0,B0,B0,B0] :: 4 Word.word))" +|" interruptType_to_bits I_S_Software = ( (vec_of_bits [B0,B0,B0,B1] :: 4 Word.word))" +|" interruptType_to_bits I_M_Software = ( (vec_of_bits [B0,B0,B1,B1] :: 4 Word.word))" +|" interruptType_to_bits I_U_Timer = ( (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word))" +|" interruptType_to_bits I_S_Timer = ( (vec_of_bits [B0,B1,B0,B1] :: 4 Word.word))" +|" interruptType_to_bits I_M_Timer = ( (vec_of_bits [B0,B1,B1,B1] :: 4 Word.word))" +|" interruptType_to_bits I_U_External = ( (vec_of_bits [B1,B0,B0,B0] :: 4 Word.word))" +|" interruptType_to_bits I_S_External = ( (vec_of_bits [B1,B0,B0,B1] :: 4 Word.word))" +|" interruptType_to_bits I_M_External = ( (vec_of_bits [B1,B0,B1,B1] :: 4 Word.word))" + + (*val ExceptionType_of_num : integer -> ExceptionType*) definition ExceptionType_of_num :: " int \<Rightarrow> ExceptionType " where @@ -693,48 +756,16 @@ fun exceptionType_to_str :: " ExceptionType \<Rightarrow> string " where |" exceptionType_to_str E_SAMO_Page_Fault = ( (''store/amo-page-fault''))" -(*val InterruptType_of_num : integer -> InterruptType*) - -definition InterruptType_of_num :: " int \<Rightarrow> InterruptType " where - " InterruptType_of_num arg0 = ( - (let p00 = arg0 in - if (((p00 = (( 0 :: int)::ii)))) then I_U_Software - else if (((p00 = (( 1 :: int)::ii)))) then I_S_Software - else if (((p00 = (( 2 :: int)::ii)))) then I_M_Software - else if (((p00 = (( 3 :: int)::ii)))) then I_U_Timer - else if (((p00 = (( 4 :: int)::ii)))) then I_S_Timer - else if (((p00 = (( 5 :: int)::ii)))) then I_M_Timer - else if (((p00 = (( 6 :: int)::ii)))) then I_U_External - else if (((p00 = (( 7 :: int)::ii)))) then I_S_External - else I_M_External))" - - -(*val num_of_InterruptType : InterruptType -> integer*) +(*val not_implemented : forall 'a. string -> M 'a*) -fun num_of_InterruptType :: " InterruptType \<Rightarrow> 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))" +definition not_implemented :: " string \<Rightarrow>((register_value),'a,(exception))monad " where + " not_implemented message = ( throw (Error_not_implemented message))" -(*val interruptType_to_bits : InterruptType -> mword ty4*) +(*val internal_error : forall 'a. string -> M 'a*) -fun interruptType_to_bits :: " InterruptType \<Rightarrow>(4)Word.word " where - " interruptType_to_bits I_U_Software = ( (vec_of_bits [B0,B0,B0,B0] :: 4 Word.word))" -|" interruptType_to_bits I_S_Software = ( (vec_of_bits [B0,B0,B0,B1] :: 4 Word.word))" -|" interruptType_to_bits I_M_Software = ( (vec_of_bits [B0,B0,B1,B1] :: 4 Word.word))" -|" interruptType_to_bits I_U_Timer = ( (vec_of_bits [B0,B1,B0,B0] :: 4 Word.word))" -|" interruptType_to_bits I_S_Timer = ( (vec_of_bits [B0,B1,B0,B1] :: 4 Word.word))" -|" interruptType_to_bits I_M_Timer = ( (vec_of_bits [B0,B1,B1,B1] :: 4 Word.word))" -|" interruptType_to_bits I_U_External = ( (vec_of_bits [B1,B0,B0,B0] :: 4 Word.word))" -|" interruptType_to_bits I_S_External = ( (vec_of_bits [B1,B0,B0,B1] :: 4 Word.word))" -|" interruptType_to_bits I_M_External = ( (vec_of_bits [B1,B0,B1,B1] :: 4 Word.word))" +definition internal_error :: " string \<Rightarrow>((register_value),'a,(exception))monad " where + " internal_error s = ( assert_exp False s \<then> throw (Error_internal_error () ))" (*val TrapVectorMode_of_num : integer -> TrapVectorMode*) @@ -765,18 +796,6 @@ definition trapVectorMode_of_bits :: "(2)Word.word \<Rightarrow> TrapVectorMode else TV_Reserved))" -(*val not_implemented : forall 'a. string -> M 'a*) - -definition not_implemented :: " string \<Rightarrow>((register_value),'a,(exception))monad " where - " not_implemented message = ( throw (Error_not_implemented message))" - - -(*val internal_error : forall 'a. string -> M 'a*) - -definition internal_error :: " string \<Rightarrow>((register_value),'a,(exception))monad " where - " internal_error s = ( assert_exp False s \<then> throw (Error_internal_error () ))" - - (*val ExtStatus_of_num : integer -> ExtStatus*) definition ExtStatus_of_num :: " int \<Rightarrow> ExtStatus " where @@ -838,17 +857,17 @@ fun num_of_SATPMode :: " SATPMode \<Rightarrow> int " where (*val satpMode_of_bits : Architecture -> mword ty4 -> maybe SATPMode*) definition satpMode_of_bits :: " Architecture \<Rightarrow>(4)Word.word \<Rightarrow>(SATPMode)option " where - " satpMode_of_bits (g__33 :: Architecture) (b__0 :: satp_mode) = ( + " satpMode_of_bits (g__19 :: Architecture) (b__0 :: satp_mode) = ( if (((b__0 = (vec_of_bits [B0,B0,B0,B0] :: 4 Word.word)))) then Some Sbare else - (case (g__33, b__0) of - (RV32, b__1) => - if (((b__1 = (vec_of_bits [B0,B0,B0,B1] :: 4 Word.word)))) then Some Sv32 - else (case (RV32, b__1) of (g__34, g__35) => None ) - | (RV64, b__2) => - if (((b__2 = (vec_of_bits [B1,B0,B0,B0] :: 4 Word.word)))) then Some Sv39 - else (case (RV64, b__2) of (g__34, g__35) => None ) - | (g__34, g__35) => None + (case (g__19, b__0) of + (RV32, b__0) => + if (((b__0 = (vec_of_bits [B0,B0,B0,B1] :: 4 Word.word)))) then Some Sv32 + else (case (RV32, b__0) of (g__20, g__21) => None ) + | (RV64, b__0) => + if (((b__0 = (vec_of_bits [B1,B0,B0,B0] :: 4 Word.word)))) then Some Sv39 + else (case (RV64, b__0) of (g__20, g__21) => None ) + | (g__20, g__21) => None ))" @@ -988,6 +1007,24 @@ fun num_of_ropw :: " ropw \<Rightarrow> int " where |" num_of_ropw RISCV_SRAW = ( (( 4 :: int)::ii))" +(*val sopw_of_num : integer -> sopw*) + +definition sopw_of_num :: " int \<Rightarrow> sopw " where + " sopw_of_num arg0 = ( + (let p00 = arg0 in + if (((p00 = (( 0 :: int)::ii)))) then RISCV_SLLIW + else if (((p00 = (( 1 :: int)::ii)))) then RISCV_SRLIW + else RISCV_SRAIW))" + + +(*val num_of_sopw : sopw -> integer*) + +fun num_of_sopw :: " sopw \<Rightarrow> 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 \<Rightarrow> amoop " where @@ -1036,93 +1073,72 @@ fun num_of_csrop :: " csrop \<Rightarrow> int " where |" num_of_csrop CSRRC = ( (( 2 :: int)::ii))" -(*val word_width_of_num : integer -> word_width*) - -definition word_width_of_num :: " int \<Rightarrow> word_width " where - " word_width_of_num arg0 = ( - (let p00 = arg0 in - if (((p00 = (( 0 :: int)::ii)))) then BYTE - else if (((p00 = (( 1 :: int)::ii)))) then HALF - else if (((p00 = (( 2 :: int)::ii)))) then WORD - else DOUBLE))" - - -(*val num_of_word_width : word_width -> integer*) - -fun num_of_word_width :: " word_width \<Rightarrow> 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 reg_name_forwards : mword ty5 -> string*) definition reg_name_forwards :: "(5)Word.word \<Rightarrow> string " where " reg_name_forwards arg0 = ( - (let p00 = arg0 in - if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) - then + (let b__0 = arg0 in + if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) then (''zero'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word)))))) then (''ra'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word)))))) then (''sp'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word)))))) then (''gp'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word)))))) then (''tp'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B1] :: 5 Word.word)))))) then (''t0'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B1,B0] :: 5 Word.word)))))) then (''t1'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B1,B1] :: 5 Word.word)))))) then (''t2'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word)))))) then (''fp'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B1] :: 5 Word.word)))))) then (''s1'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B1,B0] :: 5 Word.word)))))) then (''a0'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B1,B1] :: 5 Word.word)))))) then (''a1'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word)))))) then (''a2'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B1] :: 5 Word.word)))))) then (''a3'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B1,B0] :: 5 Word.word)))))) then (''a4'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B1,B1] :: 5 Word.word)))))) then (''a5'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word)))))) then (''a6'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B1] :: 5 Word.word)))))) then (''a7'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B1,B0] :: 5 Word.word)))))) then (''s2'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B1,B1] :: 5 Word.word)))))) then (''s3'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word)))))) then (''s4'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B1] :: 5 Word.word)))))) then (''s5'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B1,B0] :: 5 Word.word)))))) then (''s6'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B1,B1] :: 5 Word.word)))))) then (''s7'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word)))))) then (''s8'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B1] :: 5 Word.word)))))) then (''s9'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B1,B0] :: 5 Word.word)))))) then (''s10'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B1,B1] :: 5 Word.word)))))) then (''s11'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B0] :: 5 Word.word)))))) then (''t3'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B1] :: 5 Word.word)))))) then (''t4'') - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B1,B0] :: 5 Word.word)))))) then (''t5'') else (''t6'')))" @@ -1267,71 +1283,70 @@ definition reg_name_backwards :: " string \<Rightarrow>(5)Word.word " where definition reg_name_forwards_matches :: "(5)Word.word \<Rightarrow> bool " where " reg_name_forwards_matches arg0 = ( - (let p00 = arg0 in - if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) - then + (let b__0 = arg0 in + if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B1,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B1,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B1,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B1,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B1,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B1,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B1,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B1,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B1,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B1,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B1,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B1,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B1,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B1,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B1,B1] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B1,B1] :: 5 Word.word)))))) then True else False))" @@ -1424,291 +1439,291 @@ definition reg_name_backwards_matches :: " string \<Rightarrow> bool " where definition reg_name_matches_prefix :: " string \<Rightarrow>((5)Word.word*int)option " where " reg_name_matches_prefix arg0 = ( - (let stringappend_18140 = arg0 in - if (((((string_startswith stringappend_18140 (''zero''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''zero''))))) of + (let stringappend_18710 = arg0 in + if (((((string_startswith stringappend_18710 (''zero''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''zero''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''zero''))))) of + (case ((string_drop stringappend_18710 ((string_length (''zero''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''ra''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''ra''))))) of + else if (((((string_startswith stringappend_18710 (''ra''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''ra''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''ra''))))) of + (case ((string_drop stringappend_18710 ((string_length (''ra''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''sp''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''sp''))))) of + else if (((((string_startswith stringappend_18710 (''sp''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''sp''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''sp''))))) of + (case ((string_drop stringappend_18710 ((string_length (''sp''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''gp''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''gp''))))) of + else if (((((string_startswith stringappend_18710 (''gp''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''gp''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''gp''))))) of + (case ((string_drop stringappend_18710 ((string_length (''gp''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''tp''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''tp''))))) of + else if (((((string_startswith stringappend_18710 (''tp''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''tp''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''tp''))))) of + (case ((string_drop stringappend_18710 ((string_length (''tp''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''t0''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''t0''))))) of + else if (((((string_startswith stringappend_18710 (''t0''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''t0''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''t0''))))) of + (case ((string_drop stringappend_18710 ((string_length (''t0''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B0,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''t1''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''t1''))))) of + else if (((((string_startswith stringappend_18710 (''t1''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''t1''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''t1''))))) of + (case ((string_drop stringappend_18710 ((string_length (''t1''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''t2''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''t2''))))) of + else if (((((string_startswith stringappend_18710 (''t2''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''t2''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''t2''))))) of + (case ((string_drop stringappend_18710 ((string_length (''t2''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''fp''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''fp''))))) of + else if (((((string_startswith stringappend_18710 (''fp''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''fp''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''fp''))))) of + (case ((string_drop stringappend_18710 ((string_length (''fp''))))) of s0 => Some ((vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s1''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s1''))))) of + else if (((((string_startswith stringappend_18710 (''s1''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s1''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s1''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s1''))))) of s0 => Some ((vec_of_bits [B0,B1,B0,B0,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''a0''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''a0''))))) of + else if (((((string_startswith stringappend_18710 (''a0''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''a0''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''a0''))))) of + (case ((string_drop stringappend_18710 ((string_length (''a0''))))) of s0 => Some ((vec_of_bits [B0,B1,B0,B1,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''a1''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''a1''))))) of + else if (((((string_startswith stringappend_18710 (''a1''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''a1''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''a1''))))) of + (case ((string_drop stringappend_18710 ((string_length (''a1''))))) of s0 => Some ((vec_of_bits [B0,B1,B0,B1,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''a2''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''a2''))))) of + else if (((((string_startswith stringappend_18710 (''a2''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''a2''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''a2''))))) of + (case ((string_drop stringappend_18710 ((string_length (''a2''))))) of s0 => Some ((vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''a3''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''a3''))))) of + else if (((((string_startswith stringappend_18710 (''a3''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''a3''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''a3''))))) of + (case ((string_drop stringappend_18710 ((string_length (''a3''))))) of s0 => Some ((vec_of_bits [B0,B1,B1,B0,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''a4''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''a4''))))) of + else if (((((string_startswith stringappend_18710 (''a4''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''a4''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''a4''))))) of + (case ((string_drop stringappend_18710 ((string_length (''a4''))))) of s0 => Some ((vec_of_bits [B0,B1,B1,B1,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''a5''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''a5''))))) of + else if (((((string_startswith stringappend_18710 (''a5''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''a5''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''a5''))))) of + (case ((string_drop stringappend_18710 ((string_length (''a5''))))) of s0 => Some ((vec_of_bits [B0,B1,B1,B1,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''a6''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''a6''))))) of + else if (((((string_startswith stringappend_18710 (''a6''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''a6''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''a6''))))) of + (case ((string_drop stringappend_18710 ((string_length (''a6''))))) of s0 => Some ((vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''a7''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''a7''))))) of + else if (((((string_startswith stringappend_18710 (''a7''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''a7''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''a7''))))) of + (case ((string_drop stringappend_18710 ((string_length (''a7''))))) of s0 => Some ((vec_of_bits [B1,B0,B0,B0,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s2''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s2''))))) of + else if (((((string_startswith stringappend_18710 (''s2''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s2''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s2''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s2''))))) of s0 => Some ((vec_of_bits [B1,B0,B0,B1,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s3''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s3''))))) of + else if (((((string_startswith stringappend_18710 (''s3''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s3''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s3''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s3''))))) of s0 => Some ((vec_of_bits [B1,B0,B0,B1,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s4''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s4''))))) of + else if (((((string_startswith stringappend_18710 (''s4''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s4''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s4''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s4''))))) of s0 => Some ((vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s5''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s5''))))) of + else if (((((string_startswith stringappend_18710 (''s5''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s5''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s5''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s5''))))) of s0 => Some ((vec_of_bits [B1,B0,B1,B0,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s6''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s6''))))) of + else if (((((string_startswith stringappend_18710 (''s6''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s6''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s6''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s6''))))) of s0 => Some ((vec_of_bits [B1,B0,B1,B1,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s7''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s7''))))) of + else if (((((string_startswith stringappend_18710 (''s7''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s7''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s7''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s7''))))) of s0 => Some ((vec_of_bits [B1,B0,B1,B1,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s8''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s8''))))) of + else if (((((string_startswith stringappend_18710 (''s8''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s8''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s8''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s8''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s9''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s9''))))) of + else if (((((string_startswith stringappend_18710 (''s9''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s9''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s9''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s9''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B0,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s10''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s10''))))) of + else if (((((string_startswith stringappend_18710 (''s10''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s10''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s10''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s10''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B1,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''s11''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''s11''))))) of + else if (((((string_startswith stringappend_18710 (''s11''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''s11''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''s11''))))) of + (case ((string_drop stringappend_18710 ((string_length (''s11''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B1,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''t3''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''t3''))))) of + else if (((((string_startswith stringappend_18710 (''t3''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''t3''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''t3''))))) of + (case ((string_drop stringappend_18710 ((string_length (''t3''))))) of s0 => Some ((vec_of_bits [B1,B1,B1,B0,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''t4''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''t4''))))) of + else if (((((string_startswith stringappend_18710 (''t4''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''t4''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''t4''))))) of + (case ((string_drop stringappend_18710 ((string_length (''t4''))))) of s0 => Some ((vec_of_bits [B1,B1,B1,B0,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''t5''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''t5''))))) of + else if (((((string_startswith stringappend_18710 (''t5''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''t5''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''t5''))))) of + (case ((string_drop stringappend_18710 ((string_length (''t5''))))) of s0 => Some ((vec_of_bits [B1,B1,B1,B1,B0] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_18140 (''t6''))) \<and> ( - (case ((string_drop stringappend_18140 ((string_length (''t6''))))) of + else if (((((string_startswith stringappend_18710 (''t6''))) \<and> ( + (case ((string_drop stringappend_18710 ((string_length (''t6''))))) of s0 => True ))))) then - (case ((string_drop stringappend_18140 ((string_length (''t6''))))) of + (case ((string_drop stringappend_18710 ((string_length (''t6''))))) of s0 => Some ((vec_of_bits [B1,B1,B1,B1,B1] :: 5 Word.word), ((string_length arg0)) - ((string_length s0))) @@ -1728,22 +1743,22 @@ definition sep_forwards :: " unit \<Rightarrow> string " where definition sep_backwards :: " string \<Rightarrow> unit " where " sep_backwards arg0 = ( - (let stringappend_18070 = arg0 in + (let stringappend_18640 = arg0 in (case - (case ((opt_spc_matches_prefix stringappend_18070)) of - Some (stringappend_18080,stringappend_18090) => (stringappend_18080, stringappend_18090) + (case ((opt_spc_matches_prefix stringappend_18640)) of + Some (stringappend_18650,stringappend_18660) => (stringappend_18650, stringappend_18660) ) of - (_, stringappend_18090) => - (let stringappend_18100 = (string_drop stringappend_18070 - stringappend_18090) in - (let stringappend_18110 = (string_drop stringappend_18100 + (_, stringappend_18660) => + (let stringappend_18670 = (string_drop stringappend_18640 + stringappend_18660) in + (let stringappend_18680 = (string_drop stringappend_18670 ((string_length ('','')))) in (case - (case ((opt_spc_matches_prefix stringappend_18110)) of - Some (stringappend_18120,stringappend_18130) => (stringappend_18120, stringappend_18130) + (case ((opt_spc_matches_prefix stringappend_18680)) of + Some (stringappend_18690,stringappend_18700) => (stringappend_18690, stringappend_18700) ) of - (_, stringappend_18130) => - if(((string_drop stringappend_18110 stringappend_18130)) = ('''')) then + (_, stringappend_18700) => + if(((string_drop stringappend_18680 stringappend_18700)) = ('''')) then () else undefined ))) )))" @@ -1759,14 +1774,14 @@ definition sep_forwards_matches :: " unit \<Rightarrow> bool " where definition sep_backwards_matches :: " string \<Rightarrow> bool " where " sep_backwards_matches arg0 = ( - (let stringappend_18000 = arg0 in - if ((case ((opt_spc_matches_prefix stringappend_18000)) of - Some (stringappend_18010,stringappend_18020) => - (let stringappend_18030 = (string_drop stringappend_18000 stringappend_18020) in - if (((((string_startswith stringappend_18030 ('',''))) \<and> ((let stringappend_18040 = (string_drop stringappend_18030 ((string_length ('','')))) in - if ((case ((opt_spc_matches_prefix stringappend_18040)) of - Some (stringappend_18050,stringappend_18060) => - if(((string_drop stringappend_18040 stringappend_18060)) = ('''')) then + (let stringappend_18570 = arg0 in + if ((case ((opt_spc_matches_prefix stringappend_18570)) of + Some (stringappend_18580,stringappend_18590) => + (let stringappend_18600 = (string_drop stringappend_18570 stringappend_18590) in + if (((((string_startswith stringappend_18600 ('',''))) \<and> ((let stringappend_18610 = (string_drop stringappend_18600 ((string_length ('','')))) in + if ((case ((opt_spc_matches_prefix stringappend_18610)) of + Some (stringappend_18620,stringappend_18630) => + if(((string_drop stringappend_18610 stringappend_18630)) = ('''')) then True else False | None => False )) then @@ -1776,21 +1791,21 @@ definition sep_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then (case - (case ((opt_spc_matches_prefix stringappend_18000)) of - Some (stringappend_18010,stringappend_18020) => - (stringappend_18010, stringappend_18020) + (case ((opt_spc_matches_prefix stringappend_18570)) of + Some (stringappend_18580,stringappend_18590) => + (stringappend_18580, stringappend_18590) ) of - (_, stringappend_18020) => - (let stringappend_18030 = (string_drop stringappend_18000 stringappend_18020) in - (let stringappend_18040 = (string_drop stringappend_18030 + (_, stringappend_18590) => + (let stringappend_18600 = (string_drop stringappend_18570 stringappend_18590) in + (let stringappend_18610 = (string_drop stringappend_18600 ((string_length ('','')))) in (case - (case ((opt_spc_matches_prefix stringappend_18040)) of - Some (stringappend_18050,stringappend_18060) => - (stringappend_18050, stringappend_18060) + (case ((opt_spc_matches_prefix stringappend_18610)) of + Some (stringappend_18620,stringappend_18630) => + (stringappend_18620, stringappend_18630) ) of - (_, stringappend_18060) => - if(((string_drop stringappend_18040 stringappend_18060)) = ('''')) then + (_, stringappend_18630) => + if(((string_drop stringappend_18610 stringappend_18630)) = ('''')) then True else undefined ))) ) @@ -1801,14 +1816,14 @@ definition sep_backwards_matches :: " string \<Rightarrow> bool " where definition sep_matches_prefix :: " string \<Rightarrow>(unit*int)option " where " sep_matches_prefix arg0 = ( - (let stringappend_17930 = arg0 in - if ((case ((opt_spc_matches_prefix stringappend_17930)) of - Some (stringappend_17940,stringappend_17950) => - (let stringappend_17960 = (string_drop stringappend_17930 stringappend_17950) in - if (((((string_startswith stringappend_17960 ('',''))) \<and> ((let stringappend_17970 = (string_drop stringappend_17960 ((string_length ('','')))) in - if ((case ((opt_spc_matches_prefix stringappend_17970)) of - Some (stringappend_17980,stringappend_17990) => - (case ((string_drop stringappend_17970 stringappend_17990)) of + (let stringappend_18500 = arg0 in + if ((case ((opt_spc_matches_prefix stringappend_18500)) of + Some (stringappend_18510,stringappend_18520) => + (let stringappend_18530 = (string_drop stringappend_18500 stringappend_18520) in + if (((((string_startswith stringappend_18530 ('',''))) \<and> ((let stringappend_18540 = (string_drop stringappend_18530 ((string_length ('','')))) in + if ((case ((opt_spc_matches_prefix stringappend_18540)) of + Some (stringappend_18550,stringappend_18560) => + (case ((string_drop stringappend_18540 stringappend_18560)) of s0 => True ) | None => False @@ -1819,21 +1834,21 @@ definition sep_matches_prefix :: " string \<Rightarrow>(unit*int)option " wher else False) | None => False )) then (case - (case ((opt_spc_matches_prefix stringappend_17930)) of - Some (stringappend_17940,stringappend_17950) => - (stringappend_17940, stringappend_17950) + (case ((opt_spc_matches_prefix stringappend_18500)) of + Some (stringappend_18510,stringappend_18520) => + (stringappend_18510, stringappend_18520) ) of - (_, stringappend_17950) => - (let stringappend_17960 = (string_drop stringappend_17930 stringappend_17950) in - (let stringappend_17970 = (string_drop stringappend_17960 + (_, stringappend_18520) => + (let stringappend_18530 = (string_drop stringappend_18500 stringappend_18520) in + (let stringappend_18540 = (string_drop stringappend_18530 ((string_length ('','')))) in (case - (case ((opt_spc_matches_prefix stringappend_17970)) of - Some (stringappend_17980,stringappend_17990) => - (stringappend_17980, stringappend_17990) + (case ((opt_spc_matches_prefix stringappend_18540)) of + Some (stringappend_18550,stringappend_18560) => + (stringappend_18550, stringappend_18560) ) of - (_, stringappend_17990) => - (case ((string_drop stringappend_17970 stringappend_17990)) of + (_, stringappend_18560) => + (case ((string_drop stringappend_18540 stringappend_18560)) of s0 => Some (() , ((string_length arg0)) - ((string_length s0))) ) ))) @@ -1852,8 +1867,8 @@ fun bool_bits_forwards :: " bool \<Rightarrow>(1)Word.word " where definition bool_bits_backwards :: "(1)Word.word \<Rightarrow> bool " where " bool_bits_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then True else False))" @@ -1868,9 +1883,9 @@ fun bool_bits_forwards_matches :: " bool \<Rightarrow> bool " where definition bool_bits_backwards_matches :: "(1)Word.word \<Rightarrow> bool " where " bool_bits_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0] :: 1 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0] :: 1 Word.word)))) then True else False))" @@ -1885,8 +1900,8 @@ fun bool_not_bits_forwards :: " bool \<Rightarrow>(1)Word.word " where definition bool_not_bits_backwards :: "(1)Word.word \<Rightarrow> bool " where " bool_not_bits_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0] :: 1 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0] :: 1 Word.word)))) then True else False))" @@ -1901,9 +1916,9 @@ fun bool_not_bits_forwards_matches :: " bool \<Rightarrow> bool " where definition bool_not_bits_backwards_matches :: "(1)Word.word \<Rightarrow> bool " where " bool_not_bits_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0] :: 1 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0] :: 1 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then True else False))" @@ -1920,10 +1935,10 @@ fun size_bits_forwards :: " word_width \<Rightarrow>(2)Word.word " where definition size_bits_backwards :: "(2)Word.word \<Rightarrow> word_width " where " size_bits_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0] :: 2 Word.word)))) then BYTE - else if (((p00 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then HALF - else if (((p00 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then WORD + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0] :: 2 Word.word)))) then BYTE + else if (((b__0 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then HALF + else if (((b__0 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then WORD else DOUBLE))" @@ -1940,11 +1955,11 @@ fun size_bits_forwards_matches :: " word_width \<Rightarrow> bool " where definition size_bits_backwards_matches :: "(2)Word.word \<Rightarrow> bool " where " size_bits_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0] :: 2 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1] :: 2 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0] :: 2 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B1] :: 2 Word.word)))) then True else False))" @@ -1994,33 +2009,33 @@ definition size_mnemonic_backwards_matches :: " string \<Rightarrow> bool " wh definition size_mnemonic_matches_prefix :: " string \<Rightarrow>(word_width*int)option " where " size_mnemonic_matches_prefix arg0 = ( - (let stringappend_17890 = arg0 in - if (((((string_startswith stringappend_17890 (''b''))) \<and> ( - (case ((string_drop stringappend_17890 ((string_length (''b''))))) of + (let stringappend_18460 = arg0 in + if (((((string_startswith stringappend_18460 (''b''))) \<and> ( + (case ((string_drop stringappend_18460 ((string_length (''b''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17890 ((string_length (''b''))))) of + (case ((string_drop stringappend_18460 ((string_length (''b''))))) of s0 => Some (BYTE, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17890 (''h''))) \<and> ( - (case ((string_drop stringappend_17890 ((string_length (''h''))))) of + else if (((((string_startswith stringappend_18460 (''h''))) \<and> ( + (case ((string_drop stringappend_18460 ((string_length (''h''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17890 ((string_length (''h''))))) of + (case ((string_drop stringappend_18460 ((string_length (''h''))))) of s0 => Some (HALF, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17890 (''w''))) \<and> ( - (case ((string_drop stringappend_17890 ((string_length (''w''))))) of + else if (((((string_startswith stringappend_18460 (''w''))) \<and> ( + (case ((string_drop stringappend_18460 ((string_length (''w''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17890 ((string_length (''w''))))) of + (case ((string_drop stringappend_18460 ((string_length (''w''))))) of s0 => Some (WORD, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17890 (''d''))) \<and> ( - (case ((string_drop stringappend_17890 ((string_length (''d''))))) of + else if (((((string_startswith stringappend_18460 (''d''))) \<and> ( + (case ((string_drop stringappend_18460 ((string_length (''d''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17890 ((string_length (''d''))))) of + (case ((string_drop stringappend_18460 ((string_length (''d''))))) of s0 => Some (DOUBLE, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -3738,21 +3753,6 @@ definition haveFP :: " unit \<Rightarrow>((register_value),(bool),(exception))m return (((((get_Misa_D w__1 :: 1 Word.word)) = ((bool_to_bits True :: 1 Word.word))))))))" -(*val pc_alignment_mask : unit -> M (mword ty64)*) - -definition pc_alignment_mask :: " unit \<Rightarrow>((register_value),((64)Word.word),(exception))monad " where - " pc_alignment_mask _ = ( - read_reg misa_ref \<bind> (\<lambda> (w__0 :: Misa) . - return ((not_vec - ((EXTZ (( 64 :: int)::ii) - (if (((((get_Misa_C w__0 :: 1 Word.word)) = ((bool_to_bits True :: 1 Word.word))))) - then - (vec_of_bits [B0,B0] :: 2 Word.word) - else (vec_of_bits [B1,B0] :: 2 Word.word)) - :: 64 Word.word)) - :: 64 Word.word))))" - - (*val Mk_Minterrupts : mword ty64 -> Minterrupts*) definition Mk_Minterrupts :: "(64)Word.word \<Rightarrow> Minterrupts " where @@ -4950,6 +4950,21 @@ definition legalize_xepc :: "(64)Word.word \<Rightarrow>((register_value),((64) :: 64 Word.word))))" +(*val pc_alignment_mask : unit -> M (mword ty64)*) + +definition pc_alignment_mask :: " unit \<Rightarrow>((register_value),((64)Word.word),(exception))monad " where + " pc_alignment_mask _ = ( + read_reg misa_ref \<bind> (\<lambda> (w__0 :: Misa) . + return ((not_vec + ((EXTZ (( 64 :: int)::ii) + (if (((((get_Misa_C w__0 :: 1 Word.word)) = ((bool_to_bits True :: 1 Word.word))))) + then + (vec_of_bits [B0,B0] :: 2 Word.word) + else (vec_of_bits [B1,B0] :: 2 Word.word)) + :: 64 Word.word)) + :: 64 Word.word))))" + + (*val Mk_Counteren : mword ty32 -> Counteren*) definition Mk_Counteren :: "(32)Word.word \<Rightarrow> Counteren " where @@ -5143,7 +5158,7 @@ definition retire_instruction :: " unit \<Rightarrow>((register_value),(unit),( if (((((bool_to_bits w__0 :: 1 Word.word)) = ((bool_to_bits True :: 1 Word.word))))) then write_reg minstret_written_ref False else - (read_reg minstret_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (read_reg minstret_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: xlenbits) . write_reg minstret_ref ((add_vec_int w__1 (( 1 :: int)::ii) :: 64 Word.word)))))" @@ -6340,150 +6355,150 @@ definition csr_name :: "(12)Word.word \<Rightarrow> string " where definition csr_name_map_forwards :: "(12)Word.word \<Rightarrow> string " where " csr_name_map_forwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''ustatus'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then (''uie'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then (''utvec'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''uscratch'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''uepc'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''ucause'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (''utval'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then (''uip'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''fflags'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''frm'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (''fcsr'') - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''cycle'') - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''time'') - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''instret'') - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''cycleh'') - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''timeh'') - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''instreth'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''sstatus'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''sedeleg'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (''sideleg'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then (''sie'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then (''stvec'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word)))) then (''scounteren'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''sscratch'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''sepc'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''scause'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (''stval'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then (''sip'') - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''satp'') - else if (((p00 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B0,B1] :: 12 Word.word)))) then (''mvendorid'') - else if (((p00 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B0] :: 12 Word.word)))) then (''marchid'') - else if (((p00 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B1] :: 12 Word.word)))) then (''mimpid'') - else if (((p00 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B1,B0,B0] :: 12 Word.word)))) then (''mhartid'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''mstatus'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''misa'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''medeleg'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (''mideleg'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then (''mie'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then (''mtvec'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word)))) then (''mcounteren'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''mscratch'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''mepc'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''mcause'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (''mtval'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then (''mip'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''pmpcfg0'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''pmpcfg1'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''pmpcfg2'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (''pmpcfg3'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B0] :: 12 Word.word)))) then (''pmpaddr0'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B1] :: 12 Word.word)))) then (''pmpaddr1'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B0] :: 12 Word.word)))) then (''pmpaddr2'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B1] :: 12 Word.word)))) then (''pmpaddr3'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B0] :: 12 Word.word)))) then (''pmpaddr4'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B1] :: 12 Word.word)))) then (''pmpaddr5'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B0] :: 12 Word.word)))) then (''pmpaddr6'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B1] :: 12 Word.word)))) then (''pmpaddr7'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B0] :: 12 Word.word)))) then (''pmpaddr8'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B1] :: 12 Word.word)))) then (''pmpaddr9'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B0] :: 12 Word.word)))) then (''pmpaddr10'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B1] :: 12 Word.word)))) then (''pmpaddr11'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B0] :: 12 Word.word)))) then (''pmpaddr12'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B1] :: 12 Word.word)))) then (''pmpaddr13'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B0] :: 12 Word.word)))) then (''pmpaddr14'') - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B1] :: 12 Word.word)))) then (''pmpaddr15'') - else if (((p00 = (vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''mcycle'') - else if (((p00 = (vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''minstret'') - else if (((p00 = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''mcycleh'') - else if (((p00 = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''minstreth'') - else if (((p00 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (''tselect'') - else if (((p00 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word)))) then (''tdata1'') - else if (((p00 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (''tdata2'') else (''tdata3'')))" @@ -7059,151 +7074,151 @@ definition csr_name_map_backwards :: " string \<Rightarrow>(12)Word.word " whe definition csr_name_map_forwards_matches :: "(12)Word.word \<Rightarrow> bool " where " csr_name_map_forwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B1,B1] :: 12 Word.word)))) then + else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B1,B1] :: 12 Word.word)))) then True else False))" @@ -7543,660 +7558,660 @@ definition csr_name_map_backwards_matches :: " string \<Rightarrow> bool " whe definition csr_name_map_matches_prefix :: " string \<Rightarrow>((12)Word.word*int)option " where " csr_name_map_matches_prefix arg0 = ( - (let stringappend_17160 = arg0 in - if (((((string_startswith stringappend_17160 (''ustatus''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''ustatus''))))) of + (let stringappend_17730 = arg0 in + if (((((string_startswith stringappend_17730 (''ustatus''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''ustatus''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''ustatus''))))) of + (case ((string_drop stringappend_17730 ((string_length (''ustatus''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''uie''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''uie''))))) of + else if (((((string_startswith stringappend_17730 (''uie''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''uie''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''uie''))))) of + (case ((string_drop stringappend_17730 ((string_length (''uie''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''utvec''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''utvec''))))) of + else if (((((string_startswith stringappend_17730 (''utvec''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''utvec''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''utvec''))))) of + (case ((string_drop stringappend_17730 ((string_length (''utvec''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''uscratch''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''uscratch''))))) of + else if (((((string_startswith stringappend_17730 (''uscratch''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''uscratch''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''uscratch''))))) of + (case ((string_drop stringappend_17730 ((string_length (''uscratch''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''uepc''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''uepc''))))) of + else if (((((string_startswith stringappend_17730 (''uepc''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''uepc''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''uepc''))))) of + (case ((string_drop stringappend_17730 ((string_length (''uepc''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''ucause''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''ucause''))))) of + else if (((((string_startswith stringappend_17730 (''ucause''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''ucause''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''ucause''))))) of + (case ((string_drop stringappend_17730 ((string_length (''ucause''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''utval''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''utval''))))) of + else if (((((string_startswith stringappend_17730 (''utval''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''utval''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''utval''))))) of + (case ((string_drop stringappend_17730 ((string_length (''utval''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''uip''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''uip''))))) of + else if (((((string_startswith stringappend_17730 (''uip''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''uip''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''uip''))))) of + (case ((string_drop stringappend_17730 ((string_length (''uip''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''fflags''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''fflags''))))) of + else if (((((string_startswith stringappend_17730 (''fflags''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''fflags''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''fflags''))))) of + (case ((string_drop stringappend_17730 ((string_length (''fflags''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''frm''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''frm''))))) of + else if (((((string_startswith stringappend_17730 (''frm''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''frm''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''frm''))))) of + (case ((string_drop stringappend_17730 ((string_length (''frm''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''fcsr''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''fcsr''))))) of + else if (((((string_startswith stringappend_17730 (''fcsr''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''fcsr''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''fcsr''))))) of + (case ((string_drop stringappend_17730 ((string_length (''fcsr''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''cycle''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''cycle''))))) of + else if (((((string_startswith stringappend_17730 (''cycle''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''cycle''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''cycle''))))) of + (case ((string_drop stringappend_17730 ((string_length (''cycle''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''time''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''time''))))) of + else if (((((string_startswith stringappend_17730 (''time''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''time''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''time''))))) of + (case ((string_drop stringappend_17730 ((string_length (''time''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''instret''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''instret''))))) of + else if (((((string_startswith stringappend_17730 (''instret''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''instret''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''instret''))))) of + (case ((string_drop stringappend_17730 ((string_length (''instret''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''cycleh''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''cycleh''))))) of + else if (((((string_startswith stringappend_17730 (''cycleh''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''cycleh''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''cycleh''))))) of + (case ((string_drop stringappend_17730 ((string_length (''cycleh''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''timeh''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''timeh''))))) of + else if (((((string_startswith stringappend_17730 (''timeh''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''timeh''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''timeh''))))) of + (case ((string_drop stringappend_17730 ((string_length (''timeh''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''instreth''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''instreth''))))) of + else if (((((string_startswith stringappend_17730 (''instreth''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''instreth''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''instreth''))))) of + (case ((string_drop stringappend_17730 ((string_length (''instreth''))))) of s0 => Some ((vec_of_bits [B1,B1,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''sstatus''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''sstatus''))))) of + else if (((((string_startswith stringappend_17730 (''sstatus''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''sstatus''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''sstatus''))))) of + (case ((string_drop stringappend_17730 ((string_length (''sstatus''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''sedeleg''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''sedeleg''))))) of + else if (((((string_startswith stringappend_17730 (''sedeleg''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''sedeleg''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''sedeleg''))))) of + (case ((string_drop stringappend_17730 ((string_length (''sedeleg''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''sideleg''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''sideleg''))))) of + else if (((((string_startswith stringappend_17730 (''sideleg''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''sideleg''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''sideleg''))))) of + (case ((string_drop stringappend_17730 ((string_length (''sideleg''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''sie''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''sie''))))) of + else if (((((string_startswith stringappend_17730 (''sie''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''sie''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''sie''))))) of + (case ((string_drop stringappend_17730 ((string_length (''sie''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''stvec''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''stvec''))))) of + else if (((((string_startswith stringappend_17730 (''stvec''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''stvec''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''stvec''))))) of + (case ((string_drop stringappend_17730 ((string_length (''stvec''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''scounteren''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''scounteren''))))) of + else if (((((string_startswith stringappend_17730 (''scounteren''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''scounteren''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''scounteren''))))) of + (case ((string_drop stringappend_17730 ((string_length (''scounteren''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''sscratch''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''sscratch''))))) of + else if (((((string_startswith stringappend_17730 (''sscratch''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''sscratch''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''sscratch''))))) of + (case ((string_drop stringappend_17730 ((string_length (''sscratch''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''sepc''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''sepc''))))) of + else if (((((string_startswith stringappend_17730 (''sepc''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''sepc''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''sepc''))))) of + (case ((string_drop stringappend_17730 ((string_length (''sepc''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''scause''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''scause''))))) of + else if (((((string_startswith stringappend_17730 (''scause''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''scause''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''scause''))))) of + (case ((string_drop stringappend_17730 ((string_length (''scause''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''stval''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''stval''))))) of + else if (((((string_startswith stringappend_17730 (''stval''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''stval''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''stval''))))) of + (case ((string_drop stringappend_17730 ((string_length (''stval''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''sip''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''sip''))))) of + else if (((((string_startswith stringappend_17730 (''sip''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''sip''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''sip''))))) of + (case ((string_drop stringappend_17730 ((string_length (''sip''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''satp''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''satp''))))) of + else if (((((string_startswith stringappend_17730 (''satp''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''satp''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''satp''))))) of + (case ((string_drop stringappend_17730 ((string_length (''satp''))))) of s0 => Some ((vec_of_bits [B0,B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mvendorid''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mvendorid''))))) of + else if (((((string_startswith stringappend_17730 (''mvendorid''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mvendorid''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mvendorid''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mvendorid''))))) of s0 => Some ((vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''marchid''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''marchid''))))) of + else if (((((string_startswith stringappend_17730 (''marchid''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''marchid''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''marchid''))))) of + (case ((string_drop stringappend_17730 ((string_length (''marchid''))))) of s0 => Some ((vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mimpid''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mimpid''))))) of + else if (((((string_startswith stringappend_17730 (''mimpid''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mimpid''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mimpid''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mimpid''))))) of s0 => Some ((vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mhartid''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mhartid''))))) of + else if (((((string_startswith stringappend_17730 (''mhartid''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mhartid''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mhartid''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mhartid''))))) of s0 => Some ((vec_of_bits [B1,B1,B1,B1,B0,B0,B0,B1,B0,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mstatus''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mstatus''))))) of + else if (((((string_startswith stringappend_17730 (''mstatus''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mstatus''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mstatus''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mstatus''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''misa''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''misa''))))) of + else if (((((string_startswith stringappend_17730 (''misa''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''misa''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''misa''))))) of + (case ((string_drop stringappend_17730 ((string_length (''misa''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''medeleg''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''medeleg''))))) of + else if (((((string_startswith stringappend_17730 (''medeleg''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''medeleg''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''medeleg''))))) of + (case ((string_drop stringappend_17730 ((string_length (''medeleg''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mideleg''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mideleg''))))) of + else if (((((string_startswith stringappend_17730 (''mideleg''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mideleg''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mideleg''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mideleg''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mie''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mie''))))) of + else if (((((string_startswith stringappend_17730 (''mie''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mie''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mie''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mie''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mtvec''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mtvec''))))) of + else if (((((string_startswith stringappend_17730 (''mtvec''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mtvec''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mtvec''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mtvec''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mcounteren''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mcounteren''))))) of + else if (((((string_startswith stringappend_17730 (''mcounteren''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mcounteren''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mcounteren''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mcounteren''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B1,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mscratch''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mscratch''))))) of + else if (((((string_startswith stringappend_17730 (''mscratch''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mscratch''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mscratch''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mscratch''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mepc''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mepc''))))) of + else if (((((string_startswith stringappend_17730 (''mepc''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mepc''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mepc''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mepc''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mcause''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mcause''))))) of + else if (((((string_startswith stringappend_17730 (''mcause''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mcause''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mcause''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mcause''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mtval''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mtval''))))) of + else if (((((string_startswith stringappend_17730 (''mtval''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mtval''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mtval''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mtval''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mip''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mip''))))) of + else if (((((string_startswith stringappend_17730 (''mip''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mip''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mip''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mip''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpcfg0''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpcfg0''))))) of + else if (((((string_startswith stringappend_17730 (''pmpcfg0''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpcfg0''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpcfg0''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpcfg0''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpcfg1''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpcfg1''))))) of + else if (((((string_startswith stringappend_17730 (''pmpcfg1''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpcfg1''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpcfg1''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpcfg1''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpcfg2''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpcfg2''))))) of + else if (((((string_startswith stringappend_17730 (''pmpcfg2''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpcfg2''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpcfg2''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpcfg2''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpcfg3''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpcfg3''))))) of + else if (((((string_startswith stringappend_17730 (''pmpcfg3''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpcfg3''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpcfg3''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpcfg3''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr0''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr0''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr0''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr0''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr0''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr0''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr1''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr1''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr1''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr1''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr1''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr1''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr2''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr2''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr2''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr2''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr2''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr2''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr3''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr3''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr3''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr3''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr3''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr3''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr4''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr4''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr4''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr4''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr4''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr4''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr5''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr5''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr5''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr5''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr5''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr5''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr6''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr6''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr6''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr6''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr6''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr6''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr7''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr7''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr7''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr7''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr7''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr7''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B1,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr8''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr8''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr8''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr8''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr8''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr8''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr9''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr9''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr9''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr9''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr9''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr9''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr10''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr10''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr10''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr10''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr10''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr10''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr11''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr11''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr11''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr11''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr11''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr11''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr12''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr12''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr12''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr12''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr12''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr12''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr13''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr13''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr13''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr13''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr13''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr13''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr14''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr14''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr14''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr14''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr14''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr14''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''pmpaddr15''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr15''))))) of + else if (((((string_startswith stringappend_17730 (''pmpaddr15''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr15''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''pmpaddr15''))))) of + (case ((string_drop stringappend_17730 ((string_length (''pmpaddr15''))))) of s0 => Some ((vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B1,B1,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mcycle''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mcycle''))))) of + else if (((((string_startswith stringappend_17730 (''mcycle''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mcycle''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mcycle''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mcycle''))))) of s0 => Some ((vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''minstret''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''minstret''))))) of + else if (((((string_startswith stringappend_17730 (''minstret''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''minstret''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''minstret''))))) of + (case ((string_drop stringappend_17730 ((string_length (''minstret''))))) of s0 => Some ((vec_of_bits [B1,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''mcycleh''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''mcycleh''))))) of + else if (((((string_startswith stringappend_17730 (''mcycleh''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''mcycleh''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''mcycleh''))))) of + (case ((string_drop stringappend_17730 ((string_length (''mcycleh''))))) of s0 => Some ((vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''minstreth''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''minstreth''))))) of + else if (((((string_startswith stringappend_17730 (''minstreth''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''minstreth''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''minstreth''))))) of + (case ((string_drop stringappend_17730 ((string_length (''minstreth''))))) of s0 => Some ((vec_of_bits [B1,B0,B1,B1,B1,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''tselect''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''tselect''))))) of + else if (((((string_startswith stringappend_17730 (''tselect''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''tselect''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''tselect''))))) of + (case ((string_drop stringappend_17730 ((string_length (''tselect''))))) of s0 => Some ((vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''tdata1''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''tdata1''))))) of + else if (((((string_startswith stringappend_17730 (''tdata1''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''tdata1''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''tdata1''))))) of + (case ((string_drop stringappend_17730 ((string_length (''tdata1''))))) of s0 => Some ((vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''tdata2''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''tdata2''))))) of + else if (((((string_startswith stringappend_17730 (''tdata2''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''tdata2''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''tdata2''))))) of + (case ((string_drop stringappend_17730 ((string_length (''tdata2''))))) of s0 => Some ((vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B1,B0] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17160 (''tdata3''))) \<and> ( - (case ((string_drop stringappend_17160 ((string_length (''tdata3''))))) of + else if (((((string_startswith stringappend_17730 (''tdata3''))) \<and> ( + (case ((string_drop stringappend_17730 ((string_length (''tdata3''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17160 ((string_length (''tdata3''))))) of + (case ((string_drop stringappend_17730 ((string_length (''tdata3''))))) of s0 => Some ((vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B1,B1] :: 12 Word.word), ((string_length arg0)) - ((string_length s0))) @@ -8324,7 +8339,7 @@ definition check_Counteren :: "(12)Word.word \<Rightarrow> Privilege \<Rightarr return (((((get_Counteren_IR w__2 :: 1 Word.word)) = ((bool_to_bits True :: 1 Word.word)))))) else return ((case (b__0, Supervisor) of - (g__31, g__32) => + (g__17, g__18) => if (((((zopz0zIzJ_u (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word) csr)) \<and> ((zopz0zIzJ_u csr (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B1,B1,B1,B1,B1] :: 12 Word.word)))))) then @@ -8343,14 +8358,14 @@ definition check_Counteren :: "(12)Word.word \<Rightarrow> Privilege \<Rightarr return (((((get_Counteren_IR w__8 :: 1 Word.word)) = ((bool_to_bits True :: 1 Word.word)))))) else return ((case (b__3, User) of - (g__31, g__32) => + (g__17, g__18) => if (((((zopz0zIzJ_u (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word) csr)) \<and> ((zopz0zIzJ_u csr (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B1,B1,B1,B1,B1] :: 12 Word.word)))))) then False else True )) - | (g__31, g__32) => + | (g__17, g__18) => return (if (((((zopz0zIzJ_u (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1] :: 12 Word.word) csr)) \<and> ((zopz0zIzJ_u csr (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B1,B1,B1,B1,B1] :: 12 Word.word)))))) then @@ -8457,8 +8472,8 @@ definition curInterrupt :: " Privilege \<Rightarrow> Minterrupts \<Rightarrow> return (Some r)) | None => internal_error - (((op@) (''non-zero eff_mip='') - (((op@) ((string_of_bits eff_mip)) ('', but nothing pending''))))) + (((@) (''non-zero eff_mip='') + (((@) ((string_of_bits eff_mip)) ('', but nothing pending''))))) ) else if (((eff_sie \<and> (((eff_sip \<noteq> ((EXTZ (( 64 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 64 Word.word)))))))) then @@ -8468,8 +8483,8 @@ definition curInterrupt :: " Privilege \<Rightarrow> Minterrupts \<Rightarrow> return (Some r)) | None => internal_error - (((op@) (''non-zero eff_sip='') - (((op@) ((string_of_bits eff_sip)) ('', but nothing pending''))))) + (((@) (''non-zero eff_sip='') + (((@) ((string_of_bits eff_sip)) ('', but nothing pending''))))) ) else (let p = @@ -8489,9 +8504,9 @@ definition curInterrupt :: " Privilege \<Rightarrow> Minterrupts \<Rightarrow> else (''0'')) in (let (_ :: unit) = (print_endline - (((op@) ('' MTI: pend='') - (((op@) p - (((op@) ('' enbl='') (((op@) e (((op@) ('' delg='') d))))))))))) in + (((@) ('' MTI: pend='') + (((@) p + (((@) ('' enbl='') (((@) e (((@) ('' delg='') d))))))))))) in (let eff_mip = ((and_vec en_mip ((not_vec ((get_Minterrupts_bits delg :: 64 Word.word)) :: 64 Word.word)) :: 64 Word.word)) in @@ -8501,18 +8516,18 @@ definition curInterrupt :: " Privilege \<Rightarrow> Minterrupts \<Rightarrow> read_reg mstatus_ref \<bind> (\<lambda> (w__10 :: Mstatus) . (let (_ :: unit) = (print_endline - (((op@) (''mstatus='') - (((op@) ((string_of_bits ((get_Mstatus_bits w__8 :: 64 Word.word)))) - (((op@) ('' mie,sie='') - (((op@) ((string_of_bits ((get_Mstatus_MIE w__9 :: 1 Word.word)))) - (((op@) ('','') - (((op@) + (((@) (''mstatus='') + (((@) ((string_of_bits ((get_Mstatus_bits w__8 :: 64 Word.word)))) + (((@) ('' mie,sie='') + (((@) ((string_of_bits ((get_Mstatus_MIE w__9 :: 1 Word.word)))) + (((@) ('','') + (((@) ((string_of_bits ((get_Mstatus_SIE w__10 :: 1 Word.word)))) - (((op@) ('' en_mip='') - (((op@) ((string_of_bits en_mip)) - (((op@) ('' eff_mip='') - (((op@) ((string_of_bits eff_mip)) - (((op@) ('' eff_sip='') + (((@) ('' en_mip='') + (((@) ((string_of_bits en_mip)) + (((@) ('' eff_mip='') + (((@) ((string_of_bits eff_mip)) + (((@) ('' eff_sip='') ((string_of_bits eff_sip))))))))))))))))))))))))) in return None))))))))))))))))" @@ -8531,12 +8546,12 @@ definition handle_trap :: " Privilege \<Rightarrow> bool \<Rightarrow>(4)Word.w xlenbits option) = ( (let (_ :: unit) = (print_endline - (((op@) (''handling '') - (((op@) (if intr then (''int#'') else (''exc#'')) - (((op@) ((string_of_bits c)) - (((op@) ('' at priv '') - (((op@) ((privLevel_to_str del_priv)) - (((op@) ('' with tval '') + (((@) (''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 :: 64 Word.word))))))))))))))))) in (case del_priv of Machine => @@ -8554,10 +8569,10 @@ definition handle_trap :: " Privilege \<Rightarrow> bool \<Rightarrow>(4)Word.w read_reg mstatus_ref \<bind> (\<lambda> (w__3 :: Mstatus) . (let (_ :: unit) = (print_endline - (((op@) (''CSR mstatus <- '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__2 :: 64 Word.word)))) - (((op@) ('' (input: '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__3 :: 64 Word.word)))) ('')'')))))))))) in + (((@) (''CSR mstatus <- '') + (((@) ((string_of_bits ((get_Mstatus_bits w__2 :: 64 Word.word)))) + (((@) ('' (input: '') + (((@) ((string_of_bits ((get_Mstatus_bits w__3 :: 64 Word.word)))) ('')'')))))))))) in (let (_ :: unit) = (cancel_reservation () ) in read_reg mtvec_ref \<bind> (\<lambda> (w__4 :: Mtvec) . read_reg mcause_ref \<bind> (\<lambda> (w__5 :: Mcause) . @@ -8585,10 +8600,10 @@ definition handle_trap :: " Privilege \<Rightarrow> bool \<Rightarrow>(4)Word.w read_reg mstatus_ref \<bind> (\<lambda> (w__13 :: Mstatus) . (let (_ :: unit) = (print_endline - (((op@) (''CSR mstatus <- '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__12 :: 64 Word.word)))) - (((op@) ('' (input: '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__13 :: 64 Word.word)))) + (((@) (''CSR mstatus <- '') + (((@) ((string_of_bits ((get_Mstatus_bits w__12 :: 64 Word.word)))) + (((@) ('' (input: '') + (((@) ((string_of_bits ((get_Mstatus_bits w__13 :: 64 Word.word)))) ('')'')))))))))) in (let (_ :: unit) = (cancel_reservation () ) in read_reg stvec_ref \<bind> (\<lambda> (w__14 :: Mtvec) . @@ -8610,11 +8625,11 @@ definition handle_exception :: " Privilege \<Rightarrow> ctl_result \<Rightarro exception_delegatee(sync_exception_trap e) cur_priv \<bind> (\<lambda> del_priv . (let (_ :: unit) = (print_endline - (((op@) (''trapping from '') - (((op@) ((privLevel_to_str cur_priv)) - (((op@) ('' to '') - (((op@) ((privLevel_to_str del_priv)) - (((op@) ('' to handle '') + (((@) (''trapping from '') + (((@) ((privLevel_to_str cur_priv)) + (((@) ('' to '') + (((@) ((privLevel_to_str del_priv)) + (((@) ('' to handle '') ((exceptionType_to_str(sync_exception_trap e)))))))))))))) in (handle_trap del_priv False ((exceptionType_to_bits(sync_exception_trap e) :: 4 Word.word)) pc(sync_exception_excinfo e) @@ -8631,18 +8646,18 @@ definition handle_exception :: " Privilege \<Rightarrow> ctl_result \<Rightarro read_reg mstatus_ref \<bind> (\<lambda> (w__4 :: Mstatus) . (let (_ :: unit) = (print_endline - (((op@) (''CSR mstatus <- '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__3 :: 64 Word.word)))) - (((op@) ('' (input: '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__4 :: 64 Word.word)))) ('')'')))))))))) in + (((@) (''CSR mstatus <- '') + (((@) ((string_of_bits ((get_Mstatus_bits w__3 :: 64 Word.word)))) + (((@) ('' (input: '') + (((@) ((string_of_bits ((get_Mstatus_bits w__4 :: 64 Word.word)))) ('')'')))))))))) in read_reg cur_privilege_ref \<bind> (\<lambda> (w__5 :: Privilege) . (let (_ :: unit) = (print_endline - (((op@) (''ret-ing from '') - (((op@) ((privLevel_to_str prev_priv)) - (((op@) ('' to '') ((privLevel_to_str w__5))))))))) in + (((@) (''ret-ing from '') + (((@) ((privLevel_to_str prev_priv)) + (((@) ('' to '') ((privLevel_to_str w__5))))))))) in (let (_ :: unit) = (cancel_reservation () ) in - (read_reg mepc_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: 64 Word.word) . + (read_reg mepc_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: xlenbits) . (pc_alignment_mask () :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 Word.word) . return ((and_vec w__6 w__7 :: 64 Word.word))))))))))))) | (_, CTL_SRET (_)) => @@ -8662,19 +8677,19 @@ definition handle_exception :: " Privilege \<Rightarrow> ctl_result \<Rightarro read_reg mstatus_ref \<bind> (\<lambda> (w__11 :: Mstatus) . (let (_ :: unit) = (print_endline - (((op@) (''CSR mstatus <- '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__10 :: 64 Word.word)))) - (((op@) ('' (input: '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__11 :: 64 Word.word)))) + (((@) (''CSR mstatus <- '') + (((@) ((string_of_bits ((get_Mstatus_bits w__10 :: 64 Word.word)))) + (((@) ('' (input: '') + (((@) ((string_of_bits ((get_Mstatus_bits w__11 :: 64 Word.word)))) ('')'')))))))))) in read_reg cur_privilege_ref \<bind> (\<lambda> (w__12 :: Privilege) . (let (_ :: unit) = (print_endline - (((op@) (''ret-ing from '') - (((op@) ((privLevel_to_str prev_priv)) - (((op@) ('' to '') ((privLevel_to_str w__12))))))))) in + (((@) (''ret-ing from '') + (((@) ((privLevel_to_str prev_priv)) + (((@) ('' to '') ((privLevel_to_str w__12))))))))) in (let (_ :: unit) = (cancel_reservation () ) in - (read_reg sepc_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__13 :: 64 Word.word) . + (read_reg sepc_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__13 :: xlenbits) . (pc_alignment_mask () :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: 64 Word.word) . return ((and_vec w__13 w__14 :: 64 Word.word))))))))))))) ))" @@ -8686,8 +8701,8 @@ definition handle_mem_exception :: "(64)Word.word \<Rightarrow> ExceptionType \ " handle_mem_exception (addr :: xlenbits) (e :: ExceptionType) = ( (let (t :: sync_exception) = ((| sync_exception_trap = e, sync_exception_excinfo = (Some addr) |)) in read_reg cur_privilege_ref \<bind> (\<lambda> (w__0 :: Privilege) . - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . - (handle_exception w__0 (CTL_TRAP t) w__1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: xlenbits) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: xlenbits) . + (handle_exception w__0 (CTL_TRAP t) w__1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . write_reg nextPC_ref w__2)))))" @@ -8699,8 +8714,8 @@ definition handle_decode_exception :: "(64)Word.word \<Rightarrow>((register_va ((| sync_exception_trap = E_Illegal_Instr, sync_exception_excinfo = (Some instbits) |)) in read_reg cur_privilege_ref \<bind> (\<lambda> (w__0 :: Privilege) . - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . - (handle_exception w__0 (CTL_TRAP t) w__1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: xlenbits) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: xlenbits) . + (handle_exception w__0 (CTL_TRAP t) w__1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . write_reg nextPC_ref w__2)))))" @@ -8708,9 +8723,9 @@ definition handle_decode_exception :: "(64)Word.word \<Rightarrow>((register_va definition handle_interrupt :: " InterruptType \<Rightarrow> Privilege \<Rightarrow>((register_value),(unit),(exception))monad " where " handle_interrupt (i :: InterruptType) (del_priv :: Privilege) = ( - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . (handle_trap del_priv True ((interruptType_to_bits i :: 4 Word.word)) w__0 None - :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: xlenbits) . + :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . write_reg nextPC_ref w__1)))" @@ -8722,8 +8737,8 @@ definition handle_illegal :: " unit \<Rightarrow>((register_value),(unit),(exce ((| sync_exception_trap = E_Illegal_Instr, sync_exception_excinfo = None |)) in read_reg cur_privilege_ref \<bind> (\<lambda> (w__0 :: Privilege) . - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . - (handle_exception w__0 (CTL_TRAP t) w__1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: xlenbits) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: xlenbits) . + (handle_exception w__0 (CTL_TRAP t) w__1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . write_reg nextPC_ref w__2)))))" @@ -8761,10 +8776,10 @@ definition init_sys :: " unit \<Rightarrow>((register_value),(unit),(exception) write_reg minstret_written_ref False) \<then> read_reg mstatus_ref) \<bind> (\<lambda> (w__2 :: Mstatus) . return ((print_endline - (((op@) (''CSR mstatus <- '') - (((op@) ((string_of_bits ((get_Mstatus_bits w__2 :: 64 Word.word)))) - (((op@) ('' (input: '') - (((op@) + (((@) (''CSR mstatus <- '') + (((@) ((string_of_bits ((get_Mstatus_bits w__2 :: 64 Word.word)))) + (((@) ('' (input: '') + (((@) ((string_of_bits ((EXTZ (( 64 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 64 Word.word)))) ('')'')))))))))))))))" @@ -8840,7 +8855,7 @@ definition MTIME_BASE :: "(64)Word.word " where :: 64 Word.word))" -(*val clint_load : forall 'int8_times_n. Size 'int8_times_n => mword ty64 -> integer -> M (MemoryOpResult (mword 'int8_times_n))*) +(*val clint_load : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> M (MemoryOpResult (mword 'int8_times_n))*) definition clint_load :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where " clint_load addr width = ( @@ -8850,9 +8865,9 @@ definition clint_load :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((regist read_reg mip_ref \<bind> (\<lambda> (w__0 :: Minterrupts) . (let (_ :: unit) = (print_endline - (((op@) (''clint['') - (((op@) ((string_of_bits addr)) - (((op@) (''] -> '') + (((@) (''clint['') + (((@) ((string_of_bits addr)) + (((@) (''] -> '') ((string_of_bits ((get_Minterrupts_MSI w__0 :: 1 Word.word))))))))))) in read_reg mip_ref \<bind> (\<lambda> (w__1 :: Minterrupts) . return (MemValue ((zero_extend ((get_Minterrupts_MSI w__1 :: 1 Word.word)) @@ -8862,22 +8877,22 @@ definition clint_load :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((regist (read_reg mtimecmp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: xlenbits) . (let (_ :: unit) = (print_endline - (((op@) (''clint['') - (((op@) ((string_of_bits addr)) (((op@) (''] -> '') ((string_of_bits w__2))))))))) in + (((@) (''clint['') + (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits w__2))))))))) in (read_reg mtimecmp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: xlenbits) . return (MemValue ((zero_extend w__3 (( 64 :: int)::ii) :: ( 'int8_times_n::len)Word.word)))))) else if ((((((addr = MTIME_BASE))) \<and> (((width = (( 8 :: int)::ii))))))) then (read_reg mtime_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: xlenbits) . (let (_ :: unit) = (print_endline - (((op@) (''clint['') - (((op@) ((string_of_bits addr)) (((op@) (''] -> '') ((string_of_bits w__4))))))))) in + (((@) (''clint['') + (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits w__4))))))))) in (read_reg mtime_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__5 :: xlenbits) . return (MemValue ((zero_extend w__5 (( 64 :: int)::ii) :: ( 'int8_times_n::len)Word.word)))))) else (let (_ :: unit) = (print_endline - (((op@) (''clint['') (((op@) ((string_of_bits addr)) (''] -> <not-mapped>'')))))) in + (((@) (''clint['') (((@) ((string_of_bits addr)) (''] -> <not-mapped>'')))))) in return (MemException E_Load_Access_Fault))))" @@ -8886,33 +8901,33 @@ definition clint_load :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((regist definition clint_dispatch :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " clint_dispatch _ = ( (read_reg mtime_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . - (let (_ :: unit) = (print_endline (((op@) (''clint::tick mtime <- '') ((string_of_bits w__0))))) in + (let (_ :: unit) = (print_endline (((@) (''clint::tick mtime <- '') ((string_of_bits w__0))))) in (set_Minterrupts_MTI mip_ref ((bool_to_bits False :: 1 Word.word)) \<then> (read_reg mtimecmp_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: xlenbits) . - (read_reg mtime_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . + (read_reg mtime_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: xlenbits) . if ((zopz0zIzJ_u w__1 w__2)) then (read_reg mtime_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: xlenbits) . (let (_ :: unit) = - (print_endline (((op@) ('' clint timer pending at mtime '') ((string_of_bits w__3))))) in + (print_endline (((@) ('' clint timer pending at mtime '') ((string_of_bits w__3))))) in set_Minterrupts_MTI mip_ref ((bool_to_bits True :: 1 Word.word)))) else return () )))))" -(*val clint_store : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val clint_store : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -definition clint_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition clint_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " clint_store addr width data = ( (let addr = ((sub_vec addr ((plat_clint_base () :: 64 Word.word)) :: 64 Word.word)) in if ((((((addr = MSIP_BASE))) \<and> ((((((width = (( 8 :: int)::ii)))) \<or> (((width = (( 4 :: int)::ii)))))))))) then (let (_ :: unit) = (print_endline - (((op@) (''clint['') - (((op@) ((string_of_bits addr)) - (((op@) (''] <- '') - (((op@) ((string_of_bits data)) - (((op@) ('' (mip.MSI <- '') - (((op@) + (((@) (''clint['') + (((@) ((string_of_bits addr)) + (((@) (''] <- '') + (((@) ((string_of_bits data)) + (((@) ('' (mip.MSI <- '') + (((@) ((string_of_bits ((cast_unit_vec0 ((access_vec_dec data (( 0 :: int)::ii))) :: 1 Word.word)))) ('')'')))))))))))))) in @@ -8920,21 +8935,21 @@ definition clint_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_ ((bool_to_bits (((((cast_unit_vec0 ((access_vec_dec data (( 0 :: int)::ii))) :: 1 Word.word)) = (vec_of_bits [B1] :: 1 Word.word)))) :: 1 Word.word)) \<then> - clint_dispatch () ) \<then> return (MemValue () )) + clint_dispatch () ) \<then> return (MemValue True)) else if ((((((addr = MTIMECMP_BASE))) \<and> (((width = (( 8 :: int)::ii))))))) then (let (_ :: unit) = (print_endline - (((op@) (''clint['') - (((op@) ((string_of_bits addr)) - (((op@) (''] <- '') (((op@) ((string_of_bits data)) ('' (mtimecmp)'')))))))))) in + (((@) (''clint['') + (((@) ((string_of_bits addr)) + (((@) (''] <- '') (((@) ((string_of_bits data)) ('' (mtimecmp)'')))))))))) in (write_reg mtimecmp_ref ((zero_extend data (( 64 :: int)::ii) :: 64 Word.word)) \<then> - clint_dispatch () ) \<then> return (MemValue () )) + clint_dispatch () ) \<then> return (MemValue True)) else (let (_ :: unit) = (print_endline - (((op@) (''clint['') - (((op@) ((string_of_bits addr)) - (((op@) (''] <- '') (((op@) ((string_of_bits data)) ('' (<unmapped>)'')))))))))) in + (((@) (''clint['') + (((@) ((string_of_bits addr)) + (((@) (''] <- '') (((@) ((string_of_bits data)) ('' (<unmapped>)'')))))))))) in return (MemException E_SAMO_Access_Fault))))" @@ -8942,9 +8957,9 @@ definition clint_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_ definition tick_clock :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " tick_clock _ = ( - (read_reg mcycle_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg mcycle_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . (write_reg mcycle_ref ((add_vec_int w__0 (( 1 :: int)::ii) :: 64 Word.word)) \<then> - (read_reg mtime_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (read_reg mtime_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: xlenbits) . write_reg mtime_ref ((add_vec_int w__1 (( 1 :: int)::ii) :: 64 Word.word)) \<then> clint_dispatch () )))" @@ -9090,22 +9105,22 @@ definition htif_load :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((registe (read_reg htif_tohost_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . (let (_ :: unit) = (print_endline - (((op@) (''htif['') - (((op@) ((string_of_bits addr)) (((op@) (''] -> '') ((string_of_bits w__0))))))))) in + (((@) (''htif['') + (((@) ((string_of_bits addr)) (((@) (''] -> '') ((string_of_bits w__0))))))))) in if (((width = (( 8 :: int)::ii)))) then (read_reg htif_tohost_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: xlenbits) . return (MemValue ((zero_extend w__1 (( 64 :: int)::ii) :: ( 'int8_times_n::len)Word.word)))) else return (MemException E_Load_Access_Fault))))" -(*val htif_store : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val htif_store : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -definition htif_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition htif_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " htif_store addr width data = ( (let (_ :: unit) = (print_endline - (((op@) (''htif['') - (((op@) ((string_of_bits addr)) (((op@) (''] <- '') ((string_of_bits data))))))))) in + (((@) (''htif['') + (((@) ((string_of_bits addr)) (((@) (''] <- '') ((string_of_bits data))))))))) in (let (cbits :: xlenbits) = ((EXTZ (( 64 :: int)::ii) data :: 64 Word.word)) in write_reg htif_tohost_ref cbits \<then> ((let cmd = (Mk_htif_cmd cbits) in @@ -9113,7 +9128,7 @@ definition htif_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_t (if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word)))) then (let (_ :: unit) = (print_endline - (((op@) (''htif-syscall-proxy cmd: '') + (((@) (''htif-syscall-proxy cmd: '') ((string_of_bits ((get_htif_cmd_payload cmd :: 48 Word.word))))))) in if (((((cast_unit_vec0 ((access_vec_dec ((get_htif_cmd_payload cmd :: 48 Word.word)) (( 0 :: int)::ii))) :: 1 Word.word)) = (vec_of_bits [B1] :: 1 Word.word)))) then @@ -9129,7 +9144,7 @@ definition htif_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_t return (if (((b__0 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B1] :: 8 Word.word)))) then (let (_ :: unit) = (print_endline - (((op@) (''htif-term cmd: '') + (((@) (''htif-term cmd: '') ((string_of_bits ((get_htif_cmd_payload cmd :: 48 Word.word))))))) in (let b__2 = ((get_htif_cmd_cmd cmd :: 8 Word.word)) in if (((b__2 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word)))) then () @@ -9137,9 +9152,9 @@ definition htif_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_t 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 (((op@) (''Unknown term cmd: '') ((string_of_bits b__2)))))) - else print_endline (((op@) (''htif-???? cmd: '') ((string_of_bits data)))))) \<then> - return (MemValue () )))))))" + else print_endline (((@) (''Unknown term cmd: '') ((string_of_bits b__2)))))) + else print_endline (((@) (''htif-???? cmd: '') ((string_of_bits data)))))) \<then> + return (MemValue True)))))))" (*val htif_tick : unit -> M unit*) @@ -9147,7 +9162,7 @@ definition htif_store :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_t definition htif_tick :: " unit \<Rightarrow>((register_value),(unit),(exception))monad " where " htif_tick _ = ( (read_reg htif_tohost_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . - (let (_ :: unit) = (print_endline (((op@) (''htif::tick '') ((string_of_bits w__0))))) in + (let (_ :: unit) = (print_endline (((@) (''htif::tick '') ((string_of_bits w__0))))) in write_reg htif_tohost_ref ((EXTZ (( 64 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 64 Word.word)))))" @@ -9165,21 +9180,20 @@ definition within_mmio_writable :: "(64)Word.word \<Rightarrow> int \<Rightarro (((within_clint addr width)) \<or> (((((within_htif_writable addr width)) \<and> ((width \<le> (( 8 :: int)::ii))))))))" -(*val mmio_read : forall 'int8_times_n. Size 'int8_times_n => mword ty64 -> integer -> M (MemoryOpResult (mword 'int8_times_n))*) +(*val mmio_read : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> M (MemoryOpResult (mword 'int8_times_n))*) definition mmio_read :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where " mmio_read (addr :: xlenbits) (width :: int) = ( - if ((within_clint addr width)) then - (clint_load addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) + if ((within_clint addr width)) then (clint_load addr width ) else if (((((within_htif_readable addr width)) \<and> (((( 1 :: int)::ii) \<le> width))))) then - (htif_load addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) + (htif_load addr width ) else return (MemException E_Load_Access_Fault))" -(*val mmio_write : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val mmio_write : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -definition mmio_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where - " mmio_write (addr :: xlenbits) (width :: int) (data :: 'int8_times_n bits) = ( +definition mmio_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where + " mmio_write (addr :: xlenbits) (width :: int) data = ( if ((within_clint addr width)) then clint_store addr width data else if (((((within_htif_writable addr width)) \<and> ((width \<le> (( 8 :: int)::ii)))))) then htif_store addr width data @@ -9210,36 +9224,34 @@ definition is_aligned_addr :: "(64)Word.word \<Rightarrow> int \<Rightarrow> bo (((ex_int ((hardware_mod ((Word.uint addr)) width)))) = (( 0 :: int)::ii)))" -(*val phys_mem_read : forall 'int8_times_n. Size 'int8_times_n => ReadType -> mword ty64 -> integer -> bool -> bool -> bool -> M (MemoryOpResult (mword 'int8_times_n))*) +(*val phys_mem_read : forall 'int8_times_n . Size 'int8_times_n => ReadType -> mword ty64 -> integer -> bool -> bool -> bool -> M (MemoryOpResult (mword 'int8_times_n))*) definition phys_mem_read :: " ReadType \<Rightarrow>(64)Word.word \<Rightarrow> int \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where " phys_mem_read (t :: ReadType) (addr :: xlenbits) (width :: int) (aq :: bool) (rl :: bool) (res :: bool) = ( - (RISCV_read addr width aq rl res :: ( (( 'int8_times_n::len)Word.word)option) M) \<bind> (\<lambda> (w__0 :: - (( 'int8_times_n::len)Word.word)option) . + (RISCV_read addr width aq rl res ) \<bind> (\<lambda> w__0 . return ((case (t, w__0) of (Instruction, None) => MemException E_Fetch_Access_Fault | (Data, None) => MemException E_Load_Access_Fault | (_, Some (v)) => (let (_ :: unit) = (print_endline - (((op@) (''mem['') - (((op@) ((readType_to_str t)) - (((op@) ('','') - (((op@) ((string_of_bits addr)) - (((op@) (''] -> '') ((string_of_bits v))))))))))))) in + (((@) (''mem['') + (((@) ((readType_to_str t)) + (((@) ('','') + (((@) ((string_of_bits addr)) + (((@) (''] -> '') ((string_of_bits v))))))))))))) in MemValue v) ))))" -(*val checked_mem_read : forall 'int8_times_n. Size 'int8_times_n => ReadType -> mword ty64 -> integer -> M (MemoryOpResult (mword 'int8_times_n))*) +(*val checked_mem_read : forall 'int8_times_n . Size 'int8_times_n => ReadType -> mword ty64 -> integer -> M (MemoryOpResult (mword 'int8_times_n))*) definition checked_mem_read :: " ReadType \<Rightarrow>(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where " checked_mem_read (t :: ReadType) (addr :: xlenbits) (width :: int) = ( if ((((((((readType_to_str t)) = ((readType_to_str Data))))) \<and> ((within_mmio_readable addr width))))) then - (mmio_read addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) - else if ((within_phys_mem addr width)) then - (phys_mem_read t addr width False False False :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) + (mmio_read addr width ) + else if ((within_phys_mem addr width)) then (phys_mem_read t addr width False False False ) else return (MemException E_Load_Access_Fault))" @@ -9256,32 +9268,27 @@ definition checked_mem_read :: " ReadType \<Rightarrow>(64)Word.word \<Rightarr (*val MEMr_reserved_strong_acquire : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> M (MemoryOpResult (mword 'int8_times_n))*) definition MEMr0 :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where - " MEMr0 addr width = ( (checked_mem_read Data addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M))" + " MEMr0 addr width = ( (checked_mem_read Data addr width ))" definition MEMr_acquire0 :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where - " MEMr_acquire0 addr width = ( - (checked_mem_read Data addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M))" + " MEMr_acquire0 addr width = ( (checked_mem_read Data addr width ))" definition MEMr_strong_acquire0 :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where - " MEMr_strong_acquire0 addr width = ( - (checked_mem_read Data addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M))" + " MEMr_strong_acquire0 addr width = ( (checked_mem_read Data addr width ))" definition MEMr_reserved0 :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where - " MEMr_reserved0 addr width = ( - (checked_mem_read Data addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M))" + " MEMr_reserved0 addr width = ( (checked_mem_read Data addr width ))" definition MEMr_reserved_acquire0 :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where - " MEMr_reserved_acquire0 addr width = ( - (checked_mem_read Data addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M))" + " MEMr_reserved_acquire0 addr width = ( (checked_mem_read Data addr width ))" definition MEMr_reserved_strong_acquire0 :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)MemoryOpResult),(exception))monad " where - " MEMr_reserved_strong_acquire0 addr width = ( - (checked_mem_read Data addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M))" + " MEMr_reserved_strong_acquire0 addr width = ( (checked_mem_read Data addr width ))" (*val mem_read : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> bool -> bool -> bool -> M (MemoryOpResult (mword 'int8_times_n))*) @@ -9292,19 +9299,14 @@ definition mem_read :: "(64)Word.word \<Rightarrow> int \<Rightarrow> bool \<Ri return (MemException E_Load_Addr_Align) else (case (aq, rl, res) of - (False, False, False) => - (checked_mem_read Data addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) - | (True, False, False) => (MEMr_acquire0 addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) - | (False, False, True) => - (MEMr_reserved0 addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) - | (True, False, True) => - (MEMr_reserved_acquire0 addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) + (False, False, False) => (checked_mem_read Data addr width ) + | (True, False, False) => (MEMr_acquire0 addr width ) + | (False, False, True) => (MEMr_reserved0 addr width ) + | (True, False, True) => (MEMr_reserved_acquire0 addr width ) | (False, True, False) => throw (Error_not_implemented (''load.rl'')) - | (True, True, False) => - (MEMr_strong_acquire0 addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) + | (True, True, False) => (MEMr_strong_acquire0 addr width ) | (False, True, True) => throw (Error_not_implemented (''lr.rl'')) - | (True, True, True) => - (MEMr_reserved_strong_acquire0 addr width :: ( (( 'int8_times_n::len)Word.word)MemoryOpResult) M) + | (True, True, True) => (MEMr_reserved_strong_acquire0 addr width ) ))" @@ -9327,67 +9329,65 @@ definition mem_write_ea :: "(64)Word.word \<Rightarrow> int \<Rightarrow> bool ))" -(*val phys_mem_write : forall 'int8_times_n. Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val phys_mem_write : forall 'int8_times_n. Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -definition phys_mem_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where - " phys_mem_write (addr :: xlenbits) (width :: int) (data :: 'int8_times_n bits) = ( +definition phys_mem_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where + " phys_mem_write (addr :: xlenbits) (width :: int) data = ( (let (_ :: unit) = (print_endline - (((op@) (''mem['') - (((op@) ((string_of_bits addr)) (((op@) (''] <- '') ((string_of_bits data))))))))) in - RISCV_write addr width data \<bind> (\<lambda> (w__0 :: bool) . - return (if w__0 then MemValue () - else MemException E_SAMO_Access_Fault))))" + (((@) (''mem['') + (((@) ((string_of_bits addr)) (((@) (''] <- '') ((string_of_bits data))))))))) in + RISCV_write addr width data \<bind> (\<lambda> (w__0 :: bool) . return (MemValue w__0))))" -(*val checked_mem_write : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val checked_mem_write : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -definition checked_mem_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where - " checked_mem_write (addr :: xlenbits) (width :: int) (data :: 'int8_times_n bits) = ( +definition checked_mem_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where + " checked_mem_write (addr :: xlenbits) (width :: int) data = ( if ((within_mmio_writable addr width)) then mmio_write addr width data else if ((within_phys_mem addr width)) then phys_mem_write addr width data else return (MemException E_SAMO_Access_Fault))" -(*val MEMval : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val MEMval : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -(*val MEMval_release : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val MEMval_release : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -(*val MEMval_strong_release : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val MEMval_strong_release : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -(*val MEMval_conditional : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val MEMval_conditional : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -(*val MEMval_conditional_release : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val MEMval_conditional_release : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -(*val MEMval_conditional_strong_release : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult unit)*) +(*val MEMval_conditional_strong_release : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M (MemoryOpResult bool)*) -definition MEMval :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition MEMval :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " MEMval addr width data = ( checked_mem_write addr width data )" -definition MEMval_release :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition MEMval_release :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " MEMval_release addr width data = ( checked_mem_write addr width data )" -definition MEMval_strong_release :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition MEMval_strong_release :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " MEMval_strong_release addr width data = ( checked_mem_write addr width data )" -definition MEMval_conditional :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition MEMval_conditional :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " MEMval_conditional addr width data = ( checked_mem_write addr width data )" -definition MEMval_conditional_release :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition MEMval_conditional_release :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " MEMval_conditional_release addr width data = ( checked_mem_write addr width data )" -definition MEMval_conditional_strong_release :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition MEMval_conditional_strong_release :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " MEMval_conditional_strong_release addr width data = ( checked_mem_write addr width data )" -(*val mem_write_value : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> bool -> bool -> bool -> M (MemoryOpResult unit)*) +(*val mem_write_value : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> bool -> bool -> bool -> M (MemoryOpResult bool)*) -definition mem_write_value :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow>((register_value),((unit)MemoryOpResult),(exception))monad " where +definition mem_write_value :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow>((register_value),((bool)MemoryOpResult),(exception))monad " where " mem_write_value addr width value1 aq rl con = ( if ((((((rl \<or> con))) \<and> ((\<not> ((is_aligned_addr addr width))))))) then return (MemException E_SAMO_Addr_Align) @@ -9741,27 +9741,15 @@ fun ptw_error_to_str :: " PTW_Error \<Rightarrow> string " where (*val translationException : AccessType -> PTW_Error -> ExceptionType*) -definition translationException :: " AccessType \<Rightarrow> PTW_Error \<Rightarrow> ExceptionType " where - " translationException (a :: AccessType) (f :: PTW_Error) = ( - (let (e :: ExceptionType) = - ((case (a, f) of - (ReadWrite, PTW_Access) => E_SAMO_Access_Fault - | (ReadWrite, _) => E_SAMO_Page_Fault - | (Read, PTW_Access) => E_Load_Access_Fault - | (Read, _) => E_Load_Page_Fault - | (Write, PTW_Access) => E_SAMO_Access_Fault - | (Write, _) => E_SAMO_Page_Fault - | (Fetch, PTW_Access) => E_Fetch_Access_Fault - | (Fetch, _) => E_Fetch_Page_Fault - )) in - (let (_ :: unit) = - (print_endline - (((op@) (''translationException('') - (((op@) ((accessType_to_str a)) - (((op@) ('', '') - (((op@) ((ptw_error_to_str f)) - (((op@) ('') -> '') ((exceptionType_to_str e))))))))))))) in - e)))" +fun translationException :: " AccessType \<Rightarrow> PTW_Error \<Rightarrow> ExceptionType " where + " translationException (ReadWrite :: AccessType) (PTW_Access :: PTW_Error) = ( E_SAMO_Access_Fault )" +|" translationException (ReadWrite :: AccessType) (_ :: PTW_Error) = ( E_SAMO_Page_Fault )" +|" translationException (Read :: AccessType) (PTW_Access :: PTW_Error) = ( E_Load_Access_Fault )" +|" translationException (Read :: AccessType) (_ :: PTW_Error) = ( E_Load_Page_Fault )" +|" translationException (Write :: AccessType) (PTW_Access :: PTW_Error) = ( E_SAMO_Access_Fault )" +|" translationException (Write :: AccessType) (_ :: PTW_Error) = ( E_SAMO_Page_Fault )" +|" translationException (Fetch :: AccessType) (PTW_Access :: PTW_Error) = ( E_Fetch_Access_Fault )" +|" translationException (Fetch :: AccessType) (_ :: PTW_Error) = ( E_Fetch_Page_Fault )" definition SV39_LEVEL_BITS :: " int " where @@ -10117,7 +10105,7 @@ definition update_SV39_PTE_BITS :: " SV39_PTE \<Rightarrow>(8)Word.word \<Right definition curAsid64 :: " unit \<Rightarrow>((register_value),((16)Word.word),(exception))monad " where " curAsid64 _ = ( - (read_reg satp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg satp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . (let satp64 = (Mk_Satp64 w__0) in return ((get_Satp64_Asid satp64 :: 16 Word.word)))))" @@ -10126,7 +10114,7 @@ definition curAsid64 :: " unit \<Rightarrow>((register_value),((16)Word.word),( definition curPTB39 :: " unit \<Rightarrow>((register_value),((56)Word.word),(exception))monad " where " curPTB39 _ = ( - (read_reg satp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg satp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . (let satp64 = (Mk_Satp64 w__0) in return ((EXTZ (( 56 :: int)::ii) ((shiftl ((get_Satp64_PPN satp64 :: 44 Word.word)) PAGESIZE_BITS :: 44 Word.word)) @@ -10152,21 +10140,7 @@ function (sequential,domintros) walk39 :: "(39)Word.word \<Rightarrow> AccessT (phys_mem_read Data ((EXTZ (( 64 :: int)::ii) pte_addr :: 64 Word.word)) (( 8 :: int)::ii) False False False :: ( ( 64 Word.word)MemoryOpResult) M) \<bind> (\<lambda> (w__0 :: ( 64 Word.word) MemoryOpResult) . (case w__0 of - MemException (_) => - (let (_ :: unit) = - (print_endline - (((op@) (''walk39(vaddr='') - (((op@) ((string_of_bits vaddr)) - (((op@) ('' level='') - (((op@) ((stringFromInteger level)) - (((op@) ('' pt_base='') - (((op@) ((string_of_bits ptb)) - (((op@) ('' pt_ofs='') - (((op@) ((string_of_bits pt_ofs)) - (((op@) ('' pte_addr='') - (((op@) ((string_of_bits pte_addr)) - ('': invalid pte address'')))))))))))))))))))))) in - return (PTW_Failure PTW_Access)) + MemException (_) => return (PTW_Failure PTW_Access) | MemValue (v) => (let pte = (Mk_SV39_PTE v) in (let pbits = ((get_SV39_PTE_BITS pte :: 8 Word.word)) in @@ -10243,10 +10217,6 @@ definition make_TLB39_Entry :: "(16)Word.word \<Rightarrow> bool \<Rightarrow>( TLB39_Entry_age = w__0 |)))))))" -definition TLBEntries :: " int " where - " TLBEntries = ( (( 32 :: int)::ii))" - - (*val lookupTLB39 : mword ty16 -> mword ty39 -> M (maybe ((ii * TLB39_Entry)))*) definition lookupTLB39 :: "(16)Word.word \<Rightarrow>(39)Word.word \<Rightarrow>((register_value),((int*TLB39_Entry)option),(exception))monad " where @@ -10330,7 +10300,7 @@ definition translate39 :: "(39)Word.word \<Rightarrow> AccessType \<Rightarrow> ((get_PTE_Bits_bits pbits :: 8 Word.word))))|))) in (writeTLB39 idx n_ent \<then> checked_mem_write ((EXTZ (( 64 :: int)::ii)(TLB39_Entry_pteAddr ent) :: 64 Word.word)) (( 8 :: int)::ii) - ((get_SV39_PTE_bits(TLB39_Entry_pte ent) :: 64 Word.word))) \<bind> (\<lambda> (w__2 :: unit + ((get_SV39_PTE_bits(TLB39_Entry_pte ent) :: 64 Word.word))) \<bind> (\<lambda> (w__2 :: bool MemoryOpResult) . (case w__2 of MemValue (_) => return () @@ -10357,7 +10327,7 @@ definition translate39 :: "(39)Word.word \<Rightarrow> AccessType \<Rightarrow> (let (w_pte :: SV39_PTE) = (update_SV39_PTE_BITS pte ((get_PTE_Bits_bits pbits :: 8 Word.word))) in checked_mem_write ((EXTZ (( 64 :: int)::ii) pteAddr :: 64 Word.word)) (( 8 :: int)::ii) - ((get_SV39_PTE_bits w_pte :: 64 Word.word)) \<bind> (\<lambda> (w__8 :: unit MemoryOpResult) . + ((get_SV39_PTE_bits w_pte :: 64 Word.word)) \<bind> (\<lambda> (w__8 :: bool MemoryOpResult) . (case w__8 of MemValue (_) => addToTLB39 asid vAddr pAddr w_pte pteAddr level global1 \<then> @@ -10381,7 +10351,7 @@ definition translationMode :: " Privilege \<Rightarrow>((register_value),(SATPM (let arch = (architecture ((get_Mstatus_SXL w__0 :: 2 Word.word))) in (case arch of Some (RV64) => - (read_reg satp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (read_reg satp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: xlenbits) . (let (mbits :: satp_mode) = ((get_Satp64_Mode ((Mk_Satp64 w__1)) :: 4 Word.word)) in (case ((satpMode_of_bits RV64 mbits)) of Some (m) => return m @@ -10446,8 +10416,8 @@ fun encdec_uop_forwards :: " uop \<Rightarrow>(7)Word.word " where definition encdec_uop_backwards :: "(7)Word.word \<Rightarrow> uop " where " encdec_uop_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B1,B1,B0,B1,B1,B1] :: 7 Word.word)))) then RISCV_LUI + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B1,B1,B0,B1,B1,B1] :: 7 Word.word)))) then RISCV_LUI else RISCV_AUIPC))" @@ -10462,9 +10432,9 @@ fun encdec_uop_forwards_matches :: " uop \<Rightarrow> bool " where definition encdec_uop_backwards_matches :: "(7)Word.word \<Rightarrow> bool " where " encdec_uop_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B1,B1,B0,B1,B1,B1] :: 7 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1,B0,B1,B1,B1] :: 7 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B1,B1,B0,B1,B1,B1] :: 7 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B0,B1,B0,B1,B1,B1] :: 7 Word.word)))) then True else False))" @@ -10502,19 +10472,19 @@ definition utype_mnemonic_backwards_matches :: " string \<Rightarrow> bool " w definition utype_mnemonic_matches_prefix :: " string \<Rightarrow>(uop*int)option " where " utype_mnemonic_matches_prefix arg0 = ( - (let stringappend_17140 = arg0 in - if (((((string_startswith stringappend_17140 (''lui''))) \<and> ( - (case ((string_drop stringappend_17140 ((string_length (''lui''))))) of + (let stringappend_17710 = arg0 in + if (((((string_startswith stringappend_17710 (''lui''))) \<and> ( + (case ((string_drop stringappend_17710 ((string_length (''lui''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17140 ((string_length (''lui''))))) of + (case ((string_drop stringappend_17710 ((string_length (''lui''))))) of s0 => Some (RISCV_LUI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17140 (''auipc''))) \<and> ( - (case ((string_drop stringappend_17140 ((string_length (''auipc''))))) of + else if (((((string_startswith stringappend_17710 (''auipc''))) \<and> ( + (case ((string_drop stringappend_17710 ((string_length (''auipc''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17140 ((string_length (''auipc''))))) of + (case ((string_drop stringappend_17710 ((string_length (''auipc''))))) of s0 => Some (RISCV_AUIPC, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -10535,12 +10505,12 @@ fun encdec_bop_forwards :: " bop \<Rightarrow>(3)Word.word " where definition encdec_bop_backwards :: "(3)Word.word \<Rightarrow> bop " where " encdec_bop_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then RISCV_BEQ - else if (((p00 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then RISCV_BNE - else if (((p00 = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) then RISCV_BLT - else if (((p00 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then RISCV_BGE - else if (((p00 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then RISCV_BLTU + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then RISCV_BEQ + else if (((b__0 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then RISCV_BNE + else if (((b__0 = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) then RISCV_BLT + else if (((b__0 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then RISCV_BGE + else if (((b__0 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then RISCV_BLTU else RISCV_BGEU))" @@ -10559,13 +10529,13 @@ fun encdec_bop_forwards_matches :: " bop \<Rightarrow> bool " where definition encdec_bop_backwards_matches :: "(3)Word.word \<Rightarrow> bool " where " encdec_bop_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) then True else False))" @@ -10627,47 +10597,47 @@ definition btype_mnemonic_backwards_matches :: " string \<Rightarrow> bool " w definition btype_mnemonic_matches_prefix :: " string \<Rightarrow>(bop*int)option " where " btype_mnemonic_matches_prefix arg0 = ( - (let stringappend_17080 = arg0 in - if (((((string_startswith stringappend_17080 (''beq''))) \<and> ( - (case ((string_drop stringappend_17080 ((string_length (''beq''))))) of + (let stringappend_17650 = arg0 in + if (((((string_startswith stringappend_17650 (''beq''))) \<and> ( + (case ((string_drop stringappend_17650 ((string_length (''beq''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17080 ((string_length (''beq''))))) of + (case ((string_drop stringappend_17650 ((string_length (''beq''))))) of s0 => Some (RISCV_BEQ, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17080 (''bne''))) \<and> ( - (case ((string_drop stringappend_17080 ((string_length (''bne''))))) of + else if (((((string_startswith stringappend_17650 (''bne''))) \<and> ( + (case ((string_drop stringappend_17650 ((string_length (''bne''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17080 ((string_length (''bne''))))) of + (case ((string_drop stringappend_17650 ((string_length (''bne''))))) of s0 => Some (RISCV_BNE, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17080 (''blt''))) \<and> ( - (case ((string_drop stringappend_17080 ((string_length (''blt''))))) of + else if (((((string_startswith stringappend_17650 (''blt''))) \<and> ( + (case ((string_drop stringappend_17650 ((string_length (''blt''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17080 ((string_length (''blt''))))) of + (case ((string_drop stringappend_17650 ((string_length (''blt''))))) of s0 => Some (RISCV_BLT, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17080 (''bge''))) \<and> ( - (case ((string_drop stringappend_17080 ((string_length (''bge''))))) of + else if (((((string_startswith stringappend_17650 (''bge''))) \<and> ( + (case ((string_drop stringappend_17650 ((string_length (''bge''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17080 ((string_length (''bge''))))) of + (case ((string_drop stringappend_17650 ((string_length (''bge''))))) of s0 => Some (RISCV_BGE, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17080 (''bltu''))) \<and> ( - (case ((string_drop stringappend_17080 ((string_length (''bltu''))))) of + else if (((((string_startswith stringappend_17650 (''bltu''))) \<and> ( + (case ((string_drop stringappend_17650 ((string_length (''bltu''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17080 ((string_length (''bltu''))))) of + (case ((string_drop stringappend_17650 ((string_length (''bltu''))))) of s0 => Some (RISCV_BLTU, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17080 (''bgeu''))) \<and> ( - (case ((string_drop stringappend_17080 ((string_length (''bgeu''))))) of + else if (((((string_startswith stringappend_17650 (''bgeu''))) \<and> ( + (case ((string_drop stringappend_17650 ((string_length (''bgeu''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17080 ((string_length (''bgeu''))))) of + (case ((string_drop stringappend_17650 ((string_length (''bgeu''))))) of s0 => Some (RISCV_BGEU, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -10688,12 +10658,12 @@ fun encdec_iop_forwards :: " iop \<Rightarrow>(3)Word.word " where definition encdec_iop_backwards :: "(3)Word.word \<Rightarrow> iop " where " encdec_iop_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then RISCV_ADDI - else if (((p00 = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) then RISCV_SLTI - else if (((p00 = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) then RISCV_SLTIU - else if (((p00 = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) then RISCV_XORI - else if (((p00 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then RISCV_ORI + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then RISCV_ADDI + else if (((b__0 = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) then RISCV_SLTI + else if (((b__0 = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) then RISCV_SLTIU + else if (((b__0 = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) then RISCV_XORI + else if (((b__0 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then RISCV_ORI else RISCV_ANDI))" @@ -10712,13 +10682,13 @@ fun encdec_iop_forwards_matches :: " iop \<Rightarrow> bool " where definition encdec_iop_backwards_matches :: "(3)Word.word \<Rightarrow> bool " where " encdec_iop_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) then True else False))" @@ -10780,47 +10750,47 @@ definition itype_mnemonic_backwards_matches :: " string \<Rightarrow> bool " w definition itype_mnemonic_matches_prefix :: " string \<Rightarrow>(iop*int)option " where " itype_mnemonic_matches_prefix arg0 = ( - (let stringappend_17020 = arg0 in - if (((((string_startswith stringappend_17020 (''addi''))) \<and> ( - (case ((string_drop stringappend_17020 ((string_length (''addi''))))) of + (let stringappend_17590 = arg0 in + if (((((string_startswith stringappend_17590 (''addi''))) \<and> ( + (case ((string_drop stringappend_17590 ((string_length (''addi''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17020 ((string_length (''addi''))))) of + (case ((string_drop stringappend_17590 ((string_length (''addi''))))) of s0 => Some (RISCV_ADDI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17020 (''slti''))) \<and> ( - (case ((string_drop stringappend_17020 ((string_length (''slti''))))) of + else if (((((string_startswith stringappend_17590 (''slti''))) \<and> ( + (case ((string_drop stringappend_17590 ((string_length (''slti''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17020 ((string_length (''slti''))))) of + (case ((string_drop stringappend_17590 ((string_length (''slti''))))) of s0 => Some (RISCV_SLTI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17020 (''sltiu''))) \<and> ( - (case ((string_drop stringappend_17020 ((string_length (''sltiu''))))) of + else if (((((string_startswith stringappend_17590 (''sltiu''))) \<and> ( + (case ((string_drop stringappend_17590 ((string_length (''sltiu''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17020 ((string_length (''sltiu''))))) of + (case ((string_drop stringappend_17590 ((string_length (''sltiu''))))) of s0 => Some (RISCV_SLTIU, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17020 (''xori''))) \<and> ( - (case ((string_drop stringappend_17020 ((string_length (''xori''))))) of + else if (((((string_startswith stringappend_17590 (''xori''))) \<and> ( + (case ((string_drop stringappend_17590 ((string_length (''xori''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17020 ((string_length (''xori''))))) of + (case ((string_drop stringappend_17590 ((string_length (''xori''))))) of s0 => Some (RISCV_XORI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17020 (''ori''))) \<and> ( - (case ((string_drop stringappend_17020 ((string_length (''ori''))))) of + else if (((((string_startswith stringappend_17590 (''ori''))) \<and> ( + (case ((string_drop stringappend_17590 ((string_length (''ori''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17020 ((string_length (''ori''))))) of + (case ((string_drop stringappend_17590 ((string_length (''ori''))))) of s0 => Some (RISCV_ORI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_17020 (''andi''))) \<and> ( - (case ((string_drop stringappend_17020 ((string_length (''andi''))))) of + else if (((((string_startswith stringappend_17590 (''andi''))) \<and> ( + (case ((string_drop stringappend_17590 ((string_length (''andi''))))) of s0 => True ))))) then - (case ((string_drop stringappend_17020 ((string_length (''andi''))))) of + (case ((string_drop stringappend_17590 ((string_length (''andi''))))) of s0 => Some (RISCV_ANDI, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -10838,9 +10808,9 @@ fun encdec_sop_forwards :: " sop \<Rightarrow>(3)Word.word " where definition encdec_sop_backwards :: "(3)Word.word \<Rightarrow> sop " where " encdec_sop_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then RISCV_SLLI - else if (((p00 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then RISCV_SRLI + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then RISCV_SLLI + else if (((b__0 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then RISCV_SRLI else RISCV_SRAI))" @@ -10856,10 +10826,10 @@ fun encdec_sop_forwards_matches :: " sop \<Rightarrow> bool " where definition encdec_sop_backwards_matches :: "(3)Word.word \<Rightarrow> bool " where " encdec_sop_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) then True else False))" @@ -10903,26 +10873,26 @@ definition shiftiop_mnemonic_backwards_matches :: " string \<Rightarrow> bool " definition shiftiop_mnemonic_matches_prefix :: " string \<Rightarrow>(sop*int)option " where " shiftiop_mnemonic_matches_prefix arg0 = ( - (let stringappend_16990 = arg0 in - if (((((string_startswith stringappend_16990 (''slli''))) \<and> ( - (case ((string_drop stringappend_16990 ((string_length (''slli''))))) of + (let stringappend_17560 = arg0 in + if (((((string_startswith stringappend_17560 (''slli''))) \<and> ( + (case ((string_drop stringappend_17560 ((string_length (''slli''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16990 ((string_length (''slli''))))) of + (case ((string_drop stringappend_17560 ((string_length (''slli''))))) of s0 => Some (RISCV_SLLI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16990 (''srli''))) \<and> ( - (case ((string_drop stringappend_16990 ((string_length (''srli''))))) of + else if (((((string_startswith stringappend_17560 (''srli''))) \<and> ( + (case ((string_drop stringappend_17560 ((string_length (''srli''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16990 ((string_length (''srli''))))) of + (case ((string_drop stringappend_17560 ((string_length (''srli''))))) of s0 => Some (RISCV_SRLI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16990 (''srai''))) \<and> ( - (case ((string_drop stringappend_16990 ((string_length (''srai''))))) of + else if (((((string_startswith stringappend_17560 (''srai''))) \<and> ( + (case ((string_drop stringappend_17560 ((string_length (''srai''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16990 ((string_length (''srai''))))) of + (case ((string_drop stringappend_17560 ((string_length (''srai''))))) of s0 => Some (RISCV_SRAI, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -11010,75 +10980,75 @@ definition rtype_mnemonic_backwards_matches :: " string \<Rightarrow> bool " w definition rtype_mnemonic_matches_prefix :: " string \<Rightarrow>(rop*int)option " where " rtype_mnemonic_matches_prefix arg0 = ( - (let stringappend_16890 = arg0 in - if (((((string_startswith stringappend_16890 (''add''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''add''))))) of + (let stringappend_17460 = arg0 in + if (((((string_startswith stringappend_17460 (''add''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''add''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''add''))))) of + (case ((string_drop stringappend_17460 ((string_length (''add''))))) of s0 => Some (RISCV_ADD, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''sub''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''sub''))))) of + else if (((((string_startswith stringappend_17460 (''sub''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''sub''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''sub''))))) of + (case ((string_drop stringappend_17460 ((string_length (''sub''))))) of s0 => Some (RISCV_SUB, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''sll''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''sll''))))) of + else if (((((string_startswith stringappend_17460 (''sll''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''sll''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''sll''))))) of + (case ((string_drop stringappend_17460 ((string_length (''sll''))))) of s0 => Some (RISCV_SLL, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''slt''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''slt''))))) of + else if (((((string_startswith stringappend_17460 (''slt''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''slt''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''slt''))))) of + (case ((string_drop stringappend_17460 ((string_length (''slt''))))) of s0 => Some (RISCV_SLT, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''sltu''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''sltu''))))) of + else if (((((string_startswith stringappend_17460 (''sltu''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''sltu''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''sltu''))))) of + (case ((string_drop stringappend_17460 ((string_length (''sltu''))))) of s0 => Some (RISCV_SLTU, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''xor''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''xor''))))) of + else if (((((string_startswith stringappend_17460 (''xor''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''xor''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''xor''))))) of + (case ((string_drop stringappend_17460 ((string_length (''xor''))))) of s0 => Some (RISCV_XOR, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''srl''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''srl''))))) of + else if (((((string_startswith stringappend_17460 (''srl''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''srl''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''srl''))))) of + (case ((string_drop stringappend_17460 ((string_length (''srl''))))) of s0 => Some (RISCV_SRL, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''sra''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''sra''))))) of + else if (((((string_startswith stringappend_17460 (''sra''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''sra''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''sra''))))) of + (case ((string_drop stringappend_17460 ((string_length (''sra''))))) of s0 => Some (RISCV_SRA, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''or''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''or''))))) of + else if (((((string_startswith stringappend_17460 (''or''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''or''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''or''))))) of + (case ((string_drop stringappend_17460 ((string_length (''or''))))) of s0 => Some (RISCV_OR, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16890 (''and''))) \<and> ( - (case ((string_drop stringappend_16890 ((string_length (''and''))))) of + else if (((((string_startswith stringappend_17460 (''and''))) \<and> ( + (case ((string_drop stringappend_17460 ((string_length (''and''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16890 ((string_length (''and''))))) of + (case ((string_drop stringappend_17460 ((string_length (''and''))))) of s0 => Some (RISCV_AND, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -11153,19 +11123,19 @@ definition maybe_aq_backwards_matches :: " string \<Rightarrow> bool " where definition maybe_aq_matches_prefix :: " string \<Rightarrow>(bool*int)option " where " maybe_aq_matches_prefix arg0 = ( - (let stringappend_16870 = arg0 in - if (((((string_startswith stringappend_16870 (''.aq''))) \<and> ( - (case ((string_drop stringappend_16870 ((string_length (''.aq''))))) of + (let stringappend_17440 = arg0 in + if (((((string_startswith stringappend_17440 (''.aq''))) \<and> ( + (case ((string_drop stringappend_17440 ((string_length (''.aq''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16870 ((string_length (''.aq''))))) of + (case ((string_drop stringappend_17440 ((string_length (''.aq''))))) of s0 => Some (True, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16870 (''''))) \<and> ( - (case ((string_drop stringappend_16870 ((string_length (''''))))) of + else if (((((string_startswith stringappend_17440 (''''))) \<and> ( + (case ((string_drop stringappend_17440 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16870 ((string_length (''''))))) of + (case ((string_drop stringappend_17440 ((string_length (''''))))) of s0 => Some (False, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -11204,19 +11174,19 @@ definition maybe_rl_backwards_matches :: " string \<Rightarrow> bool " where definition maybe_rl_matches_prefix :: " string \<Rightarrow>(bool*int)option " where " maybe_rl_matches_prefix arg0 = ( - (let stringappend_16850 = arg0 in - if (((((string_startswith stringappend_16850 (''.rl''))) \<and> ( - (case ((string_drop stringappend_16850 ((string_length (''.rl''))))) of + (let stringappend_17420 = arg0 in + if (((((string_startswith stringappend_17420 (''.rl''))) \<and> ( + (case ((string_drop stringappend_17420 ((string_length (''.rl''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16850 ((string_length (''.rl''))))) of + (case ((string_drop stringappend_17420 ((string_length (''.rl''))))) of s0 => Some (True, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16850 (''''))) \<and> ( - (case ((string_drop stringappend_16850 ((string_length (''''))))) of + else if (((((string_startswith stringappend_17420 (''''))) \<and> ( + (case ((string_drop stringappend_17420 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16850 ((string_length (''''))))) of + (case ((string_drop stringappend_17420 ((string_length (''''))))) of s0 => Some (False, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -11255,19 +11225,19 @@ definition maybe_u_backwards_matches :: " string \<Rightarrow> bool " where definition maybe_u_matches_prefix :: " string \<Rightarrow>(bool*int)option " where " maybe_u_matches_prefix arg0 = ( - (let stringappend_16830 = arg0 in - if (((((string_startswith stringappend_16830 (''u''))) \<and> ( - (case ((string_drop stringappend_16830 ((string_length (''u''))))) of + (let stringappend_17400 = arg0 in + if (((((string_startswith stringappend_17400 (''u''))) \<and> ( + (case ((string_drop stringappend_17400 ((string_length (''u''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16830 ((string_length (''u''))))) of + (case ((string_drop stringappend_17400 ((string_length (''u''))))) of s0 => Some (True, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16830 (''''))) \<and> ( - (case ((string_drop stringappend_16830 ((string_length (''''))))) of + else if (((((string_startswith stringappend_17400 (''''))) \<and> ( + (case ((string_drop stringappend_17400 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16830 ((string_length (''''))))) of + (case ((string_drop stringappend_17400 ((string_length (''''))))) of s0 => Some (False, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -11313,26 +11283,26 @@ definition shiftw_mnemonic_backwards_matches :: " string \<Rightarrow> bool " definition shiftw_mnemonic_matches_prefix :: " string \<Rightarrow>(sop*int)option " where " shiftw_mnemonic_matches_prefix arg0 = ( - (let stringappend_16800 = arg0 in - if (((((string_startswith stringappend_16800 (''slli''))) \<and> ( - (case ((string_drop stringappend_16800 ((string_length (''slli''))))) of + (let stringappend_17370 = arg0 in + if (((((string_startswith stringappend_17370 (''slli''))) \<and> ( + (case ((string_drop stringappend_17370 ((string_length (''slli''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16800 ((string_length (''slli''))))) of + (case ((string_drop stringappend_17370 ((string_length (''slli''))))) of s0 => Some (RISCV_SLLI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16800 (''srli''))) \<and> ( - (case ((string_drop stringappend_16800 ((string_length (''srli''))))) of + else if (((((string_startswith stringappend_17370 (''srli''))) \<and> ( + (case ((string_drop stringappend_17370 ((string_length (''srli''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16800 ((string_length (''srli''))))) of + (case ((string_drop stringappend_17370 ((string_length (''srli''))))) of s0 => Some (RISCV_SRLI, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16800 (''srai''))) \<and> ( - (case ((string_drop stringappend_16800 ((string_length (''srai''))))) of + else if (((((string_startswith stringappend_17370 (''srai''))) \<and> ( + (case ((string_drop stringappend_17370 ((string_length (''srai''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16800 ((string_length (''srai''))))) of + (case ((string_drop stringappend_17370 ((string_length (''srai''))))) of s0 => Some (RISCV_SRAI, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -11390,45 +11360,110 @@ definition rtypew_mnemonic_backwards_matches :: " string \<Rightarrow> bool " definition rtypew_mnemonic_matches_prefix :: " string \<Rightarrow>(ropw*int)option " where " rtypew_mnemonic_matches_prefix arg0 = ( - (let stringappend_16750 = arg0 in - if (((((string_startswith stringappend_16750 (''addw''))) \<and> ( - (case ((string_drop stringappend_16750 ((string_length (''addw''))))) of + (let stringappend_17320 = arg0 in + if (((((string_startswith stringappend_17320 (''addw''))) \<and> ( + (case ((string_drop stringappend_17320 ((string_length (''addw''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16750 ((string_length (''addw''))))) of + (case ((string_drop stringappend_17320 ((string_length (''addw''))))) of s0 => Some (RISCV_ADDW, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16750 (''subw''))) \<and> ( - (case ((string_drop stringappend_16750 ((string_length (''subw''))))) of + else if (((((string_startswith stringappend_17320 (''subw''))) \<and> ( + (case ((string_drop stringappend_17320 ((string_length (''subw''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16750 ((string_length (''subw''))))) of + (case ((string_drop stringappend_17320 ((string_length (''subw''))))) of s0 => Some (RISCV_SUBW, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16750 (''sllw''))) \<and> ( - (case ((string_drop stringappend_16750 ((string_length (''sllw''))))) of + else if (((((string_startswith stringappend_17320 (''sllw''))) \<and> ( + (case ((string_drop stringappend_17320 ((string_length (''sllw''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16750 ((string_length (''sllw''))))) of + (case ((string_drop stringappend_17320 ((string_length (''sllw''))))) of s0 => Some (RISCV_SLLW, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16750 (''srlw''))) \<and> ( - (case ((string_drop stringappend_16750 ((string_length (''srlw''))))) of + else if (((((string_startswith stringappend_17320 (''srlw''))) \<and> ( + (case ((string_drop stringappend_17320 ((string_length (''srlw''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16750 ((string_length (''srlw''))))) of + (case ((string_drop stringappend_17320 ((string_length (''srlw''))))) of s0 => Some (RISCV_SRLW, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16750 (''sraw''))) \<and> ( - (case ((string_drop stringappend_16750 ((string_length (''sraw''))))) of + else if (((((string_startswith stringappend_17320 (''sraw''))) \<and> ( + (case ((string_drop stringappend_17320 ((string_length (''sraw''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16750 ((string_length (''sraw''))))) of + (case ((string_drop stringappend_17320 ((string_length (''sraw''))))) of s0 => Some (RISCV_SRAW, ((string_length arg0)) - ((string_length s0))) ) else None))" +(*val shiftiwop_mnemonic_forwards : sopw -> string*) + +fun shiftiwop_mnemonic_forwards :: " sopw \<Rightarrow> string " where + " shiftiwop_mnemonic_forwards RISCV_SLLIW = ( (''slliw''))" +|" shiftiwop_mnemonic_forwards RISCV_SRLIW = ( (''srliw''))" +|" shiftiwop_mnemonic_forwards RISCV_SRAIW = ( (''sraiw''))" + + +(*val shiftiwop_mnemonic_backwards : string -> sopw*) + +definition shiftiwop_mnemonic_backwards :: " string \<Rightarrow> sopw " where + " shiftiwop_mnemonic_backwards arg0 = ( + if(arg0 = (''slliw'')) then RISCV_SLLIW else + ( + if(arg0 = (''srliw'')) then RISCV_SRLIW else + (if(arg0 = (''sraiw'')) then RISCV_SRAIW else undefined)) )" + + +(*val shiftiwop_mnemonic_forwards_matches : sopw -> bool*) + +fun shiftiwop_mnemonic_forwards_matches :: " sopw \<Rightarrow> bool " where + " shiftiwop_mnemonic_forwards_matches RISCV_SLLIW = ( True )" +|" shiftiwop_mnemonic_forwards_matches RISCV_SRLIW = ( True )" +|" shiftiwop_mnemonic_forwards_matches RISCV_SRAIW = ( True )" + + +(*val shiftiwop_mnemonic_backwards_matches : string -> bool*) + +definition shiftiwop_mnemonic_backwards_matches :: " string \<Rightarrow> bool " where + " shiftiwop_mnemonic_backwards_matches arg0 = ( + if(arg0 = (''slliw'')) then True else + ( + if(arg0 = (''srliw'')) then True else + (if(arg0 = (''sraiw'')) then True else False)) )" + + +(*val shiftiwop_mnemonic_matches_prefix : string -> maybe ((sopw * ii))*) + +definition shiftiwop_mnemonic_matches_prefix :: " string \<Rightarrow>(sopw*int)option " where + " shiftiwop_mnemonic_matches_prefix arg0 = ( + (let stringappend_17290 = arg0 in + if (((((string_startswith stringappend_17290 (''slliw''))) \<and> ( + (case ((string_drop stringappend_17290 ((string_length (''slliw''))))) of + s0 => True + ))))) then + (case ((string_drop stringappend_17290 ((string_length (''slliw''))))) of + s0 => Some (RISCV_SLLIW, ((string_length arg0)) - ((string_length s0))) + ) + else if (((((string_startswith stringappend_17290 (''srliw''))) \<and> ( + (case ((string_drop stringappend_17290 ((string_length (''srliw''))))) of + s0 => True + ))))) then + (case ((string_drop stringappend_17290 ((string_length (''srliw''))))) of + s0 => Some (RISCV_SRLIW, ((string_length arg0)) - ((string_length s0))) + ) + else if (((((string_startswith stringappend_17290 (''sraiw''))) \<and> ( + (case ((string_drop stringappend_17290 ((string_length (''sraiw''))))) of + s0 => True + ))))) then + (case ((string_drop stringappend_17290 ((string_length (''sraiw''))))) of + s0 => Some (RISCV_SRAIW, ((string_length arg0)) - ((string_length s0))) + ) + else None))" + + (*val encdec_mul_op_forwards : bool -> bool -> bool -> mword ty3*) definition encdec_mul_op_forwards :: " bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow>(3)Word.word " where @@ -11446,10 +11481,10 @@ definition encdec_mul_op_forwards :: " bool \<Rightarrow> bool \<Rightarrow> bo definition encdec_mul_op_backwards :: "(3)Word.word \<Rightarrow> bool*bool*bool " where " encdec_mul_op_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then (False, True, True) - else if (((p00 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then (True, True, True) - else if (((p00 = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) then (True, True, False) + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then (False, True, True) + else if (((b__0 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then (True, True, True) + else if (((b__0 = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) then (True, True, False) else (True, False, False)))" @@ -11471,11 +11506,11 @@ definition encdec_mul_op_forwards_matches :: " bool \<Rightarrow> bool \<Righta definition encdec_mul_op_backwards_matches :: "(3)Word.word \<Rightarrow> bool " where " encdec_mul_op_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) then True else False))" @@ -11534,33 +11569,33 @@ definition mul_mnemonic_backwards_matches :: " string \<Rightarrow> bool " whe definition mul_mnemonic_matches_prefix :: " string \<Rightarrow>((bool*bool*bool)*int)option " where " mul_mnemonic_matches_prefix arg0 = ( - (let stringappend_16710 = arg0 in - if (((((string_startswith stringappend_16710 (''mul''))) \<and> ( - (case ((string_drop stringappend_16710 ((string_length (''mul''))))) of + (let stringappend_17250 = arg0 in + if (((((string_startswith stringappend_17250 (''mul''))) \<and> ( + (case ((string_drop stringappend_17250 ((string_length (''mul''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16710 ((string_length (''mul''))))) of + (case ((string_drop stringappend_17250 ((string_length (''mul''))))) of s0 => Some ((False, True, True), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16710 (''mulh''))) \<and> ( - (case ((string_drop stringappend_16710 ((string_length (''mulh''))))) of + else if (((((string_startswith stringappend_17250 (''mulh''))) \<and> ( + (case ((string_drop stringappend_17250 ((string_length (''mulh''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16710 ((string_length (''mulh''))))) of + (case ((string_drop stringappend_17250 ((string_length (''mulh''))))) of s0 => Some ((True, True, True), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16710 (''mulhsu''))) \<and> ( - (case ((string_drop stringappend_16710 ((string_length (''mulhsu''))))) of + else if (((((string_startswith stringappend_17250 (''mulhsu''))) \<and> ( + (case ((string_drop stringappend_17250 ((string_length (''mulhsu''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16710 ((string_length (''mulhsu''))))) of + (case ((string_drop stringappend_17250 ((string_length (''mulhsu''))))) of s0 => Some ((True, True, False), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16710 (''mulhu''))) \<and> ( - (case ((string_drop stringappend_16710 ((string_length (''mulhu''))))) of + else if (((((string_startswith stringappend_17250 (''mulhu''))) \<and> ( + (case ((string_drop stringappend_17250 ((string_length (''mulhu''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16710 ((string_length (''mulhu''))))) of + (case ((string_drop stringappend_17250 ((string_length (''mulhu''))))) of s0 => Some ((True, False, False), ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -11599,19 +11634,19 @@ definition maybe_not_u_backwards_matches :: " string \<Rightarrow> bool " wher definition maybe_not_u_matches_prefix :: " string \<Rightarrow>(bool*int)option " where " maybe_not_u_matches_prefix arg0 = ( - (let stringappend_16690 = arg0 in - if (((((string_startswith stringappend_16690 (''u''))) \<and> ( - (case ((string_drop stringappend_16690 ((string_length (''u''))))) of + (let stringappend_17230 = arg0 in + if (((((string_startswith stringappend_17230 (''u''))) \<and> ( + (case ((string_drop stringappend_17230 ((string_length (''u''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16690 ((string_length (''u''))))) of + (case ((string_drop stringappend_17230 ((string_length (''u''))))) of s0 => Some (False, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16690 (''''))) \<and> ( - (case ((string_drop stringappend_16690 ((string_length (''''))))) of + else if (((((string_startswith stringappend_17230 (''''))) \<and> ( + (case ((string_drop stringappend_17230 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16690 ((string_length (''''))))) of + (case ((string_drop stringappend_17230 ((string_length (''''))))) of s0 => Some (True, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -11621,8 +11656,8 @@ definition maybe_not_u_matches_prefix :: " string \<Rightarrow>(bool*int)option definition bit_maybe_r_forwards :: "(1)Word.word \<Rightarrow> string " where " bit_maybe_r_forwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then (''r'') + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then (''r'') else ('''')))" @@ -11639,9 +11674,9 @@ definition bit_maybe_r_backwards :: " string \<Rightarrow>(1)Word.word " where definition bit_maybe_r_forwards_matches :: "(1)Word.word \<Rightarrow> bool " where " bit_maybe_r_forwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0] :: 1 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0] :: 1 Word.word)))) then True else False))" @@ -11656,21 +11691,21 @@ definition bit_maybe_r_backwards_matches :: " string \<Rightarrow> bool " wher definition bit_maybe_r_matches_prefix :: " string \<Rightarrow>((1)Word.word*int)option " where " bit_maybe_r_matches_prefix arg0 = ( - (let stringappend_16670 = arg0 in - if (((((string_startswith stringappend_16670 (''r''))) \<and> ( - (case ((string_drop stringappend_16670 ((string_length (''r''))))) of + (let stringappend_17210 = arg0 in + if (((((string_startswith stringappend_17210 (''r''))) \<and> ( + (case ((string_drop stringappend_17210 ((string_length (''r''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16670 ((string_length (''r''))))) of + (case ((string_drop stringappend_17210 ((string_length (''r''))))) of s0 => Some ((vec_of_bits [B1] :: 1 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16670 (''''))) \<and> ( - (case ((string_drop stringappend_16670 ((string_length (''''))))) of + else if (((((string_startswith stringappend_17210 (''''))) \<and> ( + (case ((string_drop stringappend_17210 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16670 ((string_length (''''))))) of + (case ((string_drop stringappend_17210 ((string_length (''''))))) of s0 => Some ((vec_of_bits [B0] :: 1 Word.word), ((string_length arg0)) - ((string_length s0))) @@ -11682,8 +11717,8 @@ definition bit_maybe_r_matches_prefix :: " string \<Rightarrow>((1)Word.word*in definition bit_maybe_w_forwards :: "(1)Word.word \<Rightarrow> string " where " bit_maybe_w_forwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then (''w'') + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then (''w'') else ('''')))" @@ -11700,9 +11735,9 @@ definition bit_maybe_w_backwards :: " string \<Rightarrow>(1)Word.word " where definition bit_maybe_w_forwards_matches :: "(1)Word.word \<Rightarrow> bool " where " bit_maybe_w_forwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0] :: 1 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0] :: 1 Word.word)))) then True else False))" @@ -11717,21 +11752,21 @@ definition bit_maybe_w_backwards_matches :: " string \<Rightarrow> bool " wher definition bit_maybe_w_matches_prefix :: " string \<Rightarrow>((1)Word.word*int)option " where " bit_maybe_w_matches_prefix arg0 = ( - (let stringappend_16650 = arg0 in - if (((((string_startswith stringappend_16650 (''w''))) \<and> ( - (case ((string_drop stringappend_16650 ((string_length (''w''))))) of + (let stringappend_17190 = arg0 in + if (((((string_startswith stringappend_17190 (''w''))) \<and> ( + (case ((string_drop stringappend_17190 ((string_length (''w''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16650 ((string_length (''w''))))) of + (case ((string_drop stringappend_17190 ((string_length (''w''))))) of s0 => Some ((vec_of_bits [B1] :: 1 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16650 (''''))) \<and> ( - (case ((string_drop stringappend_16650 ((string_length (''''))))) of + else if (((((string_startswith stringappend_17190 (''''))) \<and> ( + (case ((string_drop stringappend_17190 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16650 ((string_length (''''))))) of + (case ((string_drop stringappend_17190 ((string_length (''''))))) of s0 => Some ((vec_of_bits [B0] :: 1 Word.word), ((string_length arg0)) - ((string_length s0))) @@ -11743,8 +11778,8 @@ definition bit_maybe_w_matches_prefix :: " string \<Rightarrow>((1)Word.word*in definition bit_maybe_i_forwards :: "(1)Word.word \<Rightarrow> string " where " bit_maybe_i_forwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then (''i'') + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then (''i'') else ('''')))" @@ -11761,9 +11796,9 @@ definition bit_maybe_i_backwards :: " string \<Rightarrow>(1)Word.word " where definition bit_maybe_i_forwards_matches :: "(1)Word.word \<Rightarrow> bool " where " bit_maybe_i_forwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0] :: 1 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0] :: 1 Word.word)))) then True else False))" @@ -11778,21 +11813,21 @@ definition bit_maybe_i_backwards_matches :: " string \<Rightarrow> bool " wher definition bit_maybe_i_matches_prefix :: " string \<Rightarrow>((1)Word.word*int)option " where " bit_maybe_i_matches_prefix arg0 = ( - (let stringappend_16630 = arg0 in - if (((((string_startswith stringappend_16630 (''i''))) \<and> ( - (case ((string_drop stringappend_16630 ((string_length (''i''))))) of + (let stringappend_17170 = arg0 in + if (((((string_startswith stringappend_17170 (''i''))) \<and> ( + (case ((string_drop stringappend_17170 ((string_length (''i''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16630 ((string_length (''i''))))) of + (case ((string_drop stringappend_17170 ((string_length (''i''))))) of s0 => Some ((vec_of_bits [B1] :: 1 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16630 (''''))) \<and> ( - (case ((string_drop stringappend_16630 ((string_length (''''))))) of + else if (((((string_startswith stringappend_17170 (''''))) \<and> ( + (case ((string_drop stringappend_17170 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16630 ((string_length (''''))))) of + (case ((string_drop stringappend_17170 ((string_length (''''))))) of s0 => Some ((vec_of_bits [B0] :: 1 Word.word), ((string_length arg0)) - ((string_length s0))) @@ -11804,8 +11839,8 @@ definition bit_maybe_i_matches_prefix :: " string \<Rightarrow>((1)Word.word*in definition bit_maybe_o_forwards :: "(1)Word.word \<Rightarrow> string " where " bit_maybe_o_forwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then (''o'') + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then (''o'') else ('''')))" @@ -11822,9 +11857,9 @@ definition bit_maybe_o_backwards :: " string \<Rightarrow>(1)Word.word " where definition bit_maybe_o_forwards_matches :: "(1)Word.word \<Rightarrow> bool " where " bit_maybe_o_forwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B1] :: 1 Word.word)))) then True - else if (((p00 = (vec_of_bits [B0] :: 1 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B1] :: 1 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B0] :: 1 Word.word)))) then True else False))" @@ -11839,21 +11874,21 @@ definition bit_maybe_o_backwards_matches :: " string \<Rightarrow> bool " wher definition bit_maybe_o_matches_prefix :: " string \<Rightarrow>((1)Word.word*int)option " where " bit_maybe_o_matches_prefix arg0 = ( - (let stringappend_16610 = arg0 in - if (((((string_startswith stringappend_16610 (''o''))) \<and> ( - (case ((string_drop stringappend_16610 ((string_length (''o''))))) of + (let stringappend_17150 = arg0 in + if (((((string_startswith stringappend_17150 (''o''))) \<and> ( + (case ((string_drop stringappend_17150 ((string_length (''o''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16610 ((string_length (''o''))))) of + (case ((string_drop stringappend_17150 ((string_length (''o''))))) of s0 => Some ((vec_of_bits [B1] :: 1 Word.word), ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16610 (''''))) \<and> ( - (case ((string_drop stringappend_16610 ((string_length (''''))))) of + else if (((((string_startswith stringappend_17150 (''''))) \<and> ( + (case ((string_drop stringappend_17150 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16610 ((string_length (''''))))) of + (case ((string_drop stringappend_17150 ((string_length (''''))))) of s0 => Some ((vec_of_bits [B0] :: 1 Word.word), ((string_length arg0)) - ((string_length s0))) @@ -11865,42 +11900,43 @@ definition bit_maybe_o_matches_prefix :: " string \<Rightarrow>((1)Word.word*in definition fence_bits_forwards :: "(4)Word.word \<Rightarrow> string " where " fence_bits_forwards v__0 = ( - (let (r :: 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 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in - (let (i :: 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 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in - string_append ((bit_maybe_r_forwards r)) - ((string_append ((bit_maybe_w_forwards w)) - ((string_append ((bit_maybe_i_forwards i)) - ((string_append ((bit_maybe_o_forwards o1)) (''''))))))))))))" + (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 + string_append ((bit_maybe_i_forwards i)) + ((string_append ((bit_maybe_o_forwards o1)) + ((string_append ((bit_maybe_r_forwards r)) + ((string_append ((bit_maybe_w_forwards w)) ('''')))))))))))))" (*val fence_bits_backwards : string -> mword ty4*) definition fence_bits_backwards :: " string \<Rightarrow>(4)Word.word " where " fence_bits_backwards arg0 = ( - (let stringappend_16490 = arg0 in - (let (r, stringappend_16510) = - ((case ((bit_maybe_r_matches_prefix stringappend_16490 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16500,stringappend_16510) => (stringappend_16500, stringappend_16510) + (let stringappend_17030 = arg0 in + (let (i, stringappend_17050) = + ((case ((bit_maybe_i_matches_prefix stringappend_17030 :: (( 1 Word.word * ii))option)) of + Some (stringappend_17040,stringappend_17050) => (stringappend_17040, stringappend_17050) )) in - (let stringappend_16520 = (string_drop stringappend_16490 stringappend_16510) in - (let (w, stringappend_16540) = - ((case ((bit_maybe_w_matches_prefix stringappend_16520 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16530,stringappend_16540) => (stringappend_16530, stringappend_16540) + (let stringappend_17060 = (string_drop stringappend_17030 stringappend_17050) in + (let (o1, stringappend_17080) = + ((case ((bit_maybe_o_matches_prefix stringappend_17060 :: (( 1 Word.word * ii))option)) of + Some (stringappend_17070,stringappend_17080) => (stringappend_17070, stringappend_17080) )) in - (let stringappend_16550 = (string_drop stringappend_16520 stringappend_16540) in - (let (i, stringappend_16570) = - ((case ((bit_maybe_i_matches_prefix stringappend_16550 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16560,stringappend_16570) => (stringappend_16560, stringappend_16570) + (let stringappend_17090 = (string_drop stringappend_17060 stringappend_17080) in + (let (r, stringappend_17110) = + ((case ((bit_maybe_r_matches_prefix stringappend_17090 :: (( 1 Word.word * ii))option)) of + Some (stringappend_17100,stringappend_17110) => (stringappend_17100, stringappend_17110) )) in - (let stringappend_16580 = (string_drop stringappend_16550 stringappend_16570) in - (let (o1, stringappend_16600) = - ((case ((bit_maybe_o_matches_prefix stringappend_16580 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16590,stringappend_16600) => (stringappend_16590, stringappend_16600) - )) in if(((string_drop stringappend_16580 stringappend_16600)) = ('''')) then - ((concat_vec r - ((concat_vec w ((concat_vec i o1 :: 2 Word.word)) :: 3 Word.word)) :: 4 Word.word)) + (let stringappend_17120 = (string_drop stringappend_17090 stringappend_17110) in + (let (w, stringappend_17140) = + ((case ((bit_maybe_w_matches_prefix stringappend_17120 :: (( 1 Word.word * ii))option)) of + Some (stringappend_17130,stringappend_17140) => (stringappend_17130, stringappend_17140) + )) in if(((string_drop stringappend_17120 stringappend_17140)) = ('''')) then + ((concat_vec i + ((concat_vec o1 ((concat_vec r w :: 2 Word.word)) :: 3 Word.word)) :: 4 Word.word)) else undefined)))))))))" @@ -11914,20 +11950,20 @@ definition fence_bits_forwards_matches :: "(4)Word.word \<Rightarrow> bool " w definition fence_bits_backwards_matches :: " string \<Rightarrow> bool " where " fence_bits_backwards_matches arg0 = ( - (let stringappend_16370 = arg0 in - if ((case ((bit_maybe_r_matches_prefix stringappend_16370 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16380,stringappend_16390) => - (let stringappend_16400 = (string_drop stringappend_16370 stringappend_16390) in - if ((case ((bit_maybe_w_matches_prefix stringappend_16400 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16410,stringappend_16420) => - (let stringappend_16430 = (string_drop stringappend_16400 stringappend_16420) in - if ((case ((bit_maybe_i_matches_prefix stringappend_16430 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16440,stringappend_16450) => - (let stringappend_16460 = (string_drop stringappend_16430 stringappend_16450) in - if ((case ((bit_maybe_o_matches_prefix stringappend_16460 + (let stringappend_16910 = arg0 in + if ((case ((bit_maybe_i_matches_prefix stringappend_16910 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16920,stringappend_16930) => + (let stringappend_16940 = (string_drop stringappend_16910 stringappend_16930) in + if ((case ((bit_maybe_o_matches_prefix stringappend_16940 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16950,stringappend_16960) => + (let stringappend_16970 = (string_drop stringappend_16940 stringappend_16960) in + if ((case ((bit_maybe_r_matches_prefix stringappend_16970 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16980,stringappend_16990) => + (let stringappend_17000 = (string_drop stringappend_16970 stringappend_16990) in + if ((case ((bit_maybe_w_matches_prefix stringappend_17000 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16470,stringappend_16480) => - if(((string_drop stringappend_16460 stringappend_16480)) = ('''')) then + Some (stringappend_17010,stringappend_17020) => + if(((string_drop stringappend_17000 stringappend_17020)) = ('''')) then True else False | None => False )) then @@ -11943,29 +11979,29 @@ definition fence_bits_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let (r, stringappend_16390) = - ((case ((bit_maybe_r_matches_prefix stringappend_16370 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16380,stringappend_16390) => - (stringappend_16380, stringappend_16390) + (let (i, stringappend_16930) = + ((case ((bit_maybe_i_matches_prefix stringappend_16910 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16920,stringappend_16930) => + (stringappend_16920, stringappend_16930) )) in - (let stringappend_16400 = (string_drop stringappend_16370 stringappend_16390) in - (let (w, stringappend_16420) = - ((case ((bit_maybe_w_matches_prefix stringappend_16400 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16410,stringappend_16420) => - (stringappend_16410, stringappend_16420) + (let stringappend_16940 = (string_drop stringappend_16910 stringappend_16930) in + (let (o1, stringappend_16960) = + ((case ((bit_maybe_o_matches_prefix stringappend_16940 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16950,stringappend_16960) => + (stringappend_16950, stringappend_16960) )) in - (let stringappend_16430 = (string_drop stringappend_16400 stringappend_16420) in - (let (i, stringappend_16450) = - ((case ((bit_maybe_i_matches_prefix stringappend_16430 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16440,stringappend_16450) => - (stringappend_16440, stringappend_16450) + (let stringappend_16970 = (string_drop stringappend_16940 stringappend_16960) in + (let (r, stringappend_16990) = + ((case ((bit_maybe_r_matches_prefix stringappend_16970 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16980,stringappend_16990) => + (stringappend_16980, stringappend_16990) )) in - (let stringappend_16460 = (string_drop stringappend_16430 stringappend_16450) in - (let (o1, stringappend_16480) = - ((case ((bit_maybe_o_matches_prefix stringappend_16460 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16470,stringappend_16480) => - (stringappend_16470, stringappend_16480) - )) in if(((string_drop stringappend_16460 stringappend_16480)) = ('''')) then + (let stringappend_17000 = (string_drop stringappend_16970 stringappend_16990) in + (let (w, stringappend_17020) = + ((case ((bit_maybe_w_matches_prefix stringappend_17000 :: (( 1 Word.word * ii))option)) of + Some (stringappend_17010,stringappend_17020) => + (stringappend_17010, stringappend_17020) + )) in if(((string_drop stringappend_17000 stringappend_17020)) = ('''')) then True else undefined))))))) else False))" @@ -11974,20 +12010,20 @@ definition fence_bits_backwards_matches :: " string \<Rightarrow> bool " where definition fence_bits_matches_prefix :: " string \<Rightarrow>((4)Word.word*int)option " where " fence_bits_matches_prefix arg0 = ( - (let stringappend_16250 = arg0 in - if ((case ((bit_maybe_r_matches_prefix stringappend_16250 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16260,stringappend_16270) => - (let stringappend_16280 = (string_drop stringappend_16250 stringappend_16270) in - if ((case ((bit_maybe_w_matches_prefix stringappend_16280 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16290,stringappend_16300) => - (let stringappend_16310 = (string_drop stringappend_16280 stringappend_16300) in - if ((case ((bit_maybe_i_matches_prefix stringappend_16310 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16320,stringappend_16330) => - (let stringappend_16340 = (string_drop stringappend_16310 stringappend_16330) in - if ((case ((bit_maybe_o_matches_prefix stringappend_16340 + (let stringappend_16790 = arg0 in + if ((case ((bit_maybe_i_matches_prefix stringappend_16790 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16800,stringappend_16810) => + (let stringappend_16820 = (string_drop stringappend_16790 stringappend_16810) in + if ((case ((bit_maybe_o_matches_prefix stringappend_16820 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16830,stringappend_16840) => + (let stringappend_16850 = (string_drop stringappend_16820 stringappend_16840) in + if ((case ((bit_maybe_r_matches_prefix stringappend_16850 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16860,stringappend_16870) => + (let stringappend_16880 = (string_drop stringappend_16850 stringappend_16870) in + if ((case ((bit_maybe_w_matches_prefix stringappend_16880 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16350,stringappend_16360) => - (case ((string_drop stringappend_16340 stringappend_16360)) of + Some (stringappend_16890,stringappend_16900) => + (case ((string_drop stringappend_16880 stringappend_16900)) of s0 => True ) | None => False @@ -12004,32 +12040,32 @@ definition fence_bits_matches_prefix :: " string \<Rightarrow>((4)Word.word*int else False) | None => False )) then - (let (r, stringappend_16270) = - ((case ((bit_maybe_r_matches_prefix stringappend_16250 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16260,stringappend_16270) => - (stringappend_16260, stringappend_16270) + (let (i, stringappend_16810) = + ((case ((bit_maybe_i_matches_prefix stringappend_16790 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16800,stringappend_16810) => + (stringappend_16800, stringappend_16810) )) in - (let stringappend_16280 = (string_drop stringappend_16250 stringappend_16270) in - (let (w, stringappend_16300) = - ((case ((bit_maybe_w_matches_prefix stringappend_16280 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16290,stringappend_16300) => - (stringappend_16290, stringappend_16300) + (let stringappend_16820 = (string_drop stringappend_16790 stringappend_16810) in + (let (o1, stringappend_16840) = + ((case ((bit_maybe_o_matches_prefix stringappend_16820 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16830,stringappend_16840) => + (stringappend_16830, stringappend_16840) )) in - (let stringappend_16310 = (string_drop stringappend_16280 stringappend_16300) in - (let (i, stringappend_16330) = - ((case ((bit_maybe_i_matches_prefix stringappend_16310 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16320,stringappend_16330) => - (stringappend_16320, stringappend_16330) + (let stringappend_16850 = (string_drop stringappend_16820 stringappend_16840) in + (let (r, stringappend_16870) = + ((case ((bit_maybe_r_matches_prefix stringappend_16850 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16860,stringappend_16870) => + (stringappend_16860, stringappend_16870) )) in - (let stringappend_16340 = (string_drop stringappend_16310 stringappend_16330) in - (let (o1, stringappend_16360) = - ((case ((bit_maybe_o_matches_prefix stringappend_16340 :: (( 1 Word.word * ii))option)) of - Some (stringappend_16350,stringappend_16360) => - (stringappend_16350, stringappend_16360) + (let stringappend_16880 = (string_drop stringappend_16850 stringappend_16870) in + (let (w, stringappend_16900) = + ((case ((bit_maybe_w_matches_prefix stringappend_16880 :: (( 1 Word.word * ii))option)) of + Some (stringappend_16890,stringappend_16900) => + (stringappend_16890, stringappend_16900) )) in - (case ((string_drop stringappend_16340 stringappend_16360)) of + (case ((string_drop stringappend_16880 stringappend_16900)) of s0 => - Some ((concat_vec r ((concat_vec w ((concat_vec i o1 :: 2 Word.word)) :: 3 Word.word)) + Some ((concat_vec i ((concat_vec o1 ((concat_vec r w :: 2 Word.word)) :: 3 Word.word)) :: 4 Word.word), ((string_length arg0)) - ((string_length s0))) )))))))) @@ -12084,23 +12120,22 @@ fun encdec_amoop_forwards :: " amoop \<Rightarrow>(5)Word.word " where definition encdec_amoop_backwards :: "(5)Word.word \<Rightarrow> amoop " where " encdec_amoop_backwards arg0 = ( - (let p00 = arg0 in - if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word)))))) - then + (let b__0 = arg0 in + if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word)))))) then AMOSWAP - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) then AMOADD - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word)))))) then AMOXOR - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word)))))) then AMOAND - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word)))))) then AMOOR - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word)))))) then AMOMIN - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word)))))) then AMOMAX - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word)))))) then AMOMINU else AMOMAXU))" @@ -12123,25 +12158,24 @@ fun encdec_amoop_forwards_matches :: " amoop \<Rightarrow> bool " where definition encdec_amoop_backwards_matches :: "(5)Word.word \<Rightarrow> bool " where " encdec_amoop_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word)))))) - then + (let b__0 = arg0 in + if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B1] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B0,B1,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B1,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B0,B1,B0,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B0,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B0,B1,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B0,B0,B0] :: 5 Word.word)))))) then True - else if (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B0] :: 5 Word.word)))))) then + else if (((((regbits_to_regno b__0)) = ((regbits_to_regno (vec_of_bits [B1,B1,B1,B0,B0] :: 5 Word.word)))))) then True else False))" @@ -12222,68 +12256,68 @@ definition amo_mnemonic_backwards_matches :: " string \<Rightarrow> bool " whe definition amo_mnemonic_matches_prefix :: " string \<Rightarrow>(amoop*int)option " where " amo_mnemonic_matches_prefix arg0 = ( - (let stringappend_16160 = arg0 in - if (((((string_startswith stringappend_16160 (''amoswap''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amoswap''))))) of + (let stringappend_16700 = arg0 in + if (((((string_startswith stringappend_16700 (''amoswap''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amoswap''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amoswap''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amoswap''))))) of s0 => Some (AMOSWAP, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16160 (''amoadd''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amoadd''))))) of + else if (((((string_startswith stringappend_16700 (''amoadd''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amoadd''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amoadd''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amoadd''))))) of s0 => Some (AMOADD, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16160 (''amoxor''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amoxor''))))) of + else if (((((string_startswith stringappend_16700 (''amoxor''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amoxor''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amoxor''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amoxor''))))) of s0 => Some (AMOXOR, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16160 (''amoand''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amoand''))))) of + else if (((((string_startswith stringappend_16700 (''amoand''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amoand''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amoand''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amoand''))))) of s0 => Some (AMOAND, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16160 (''amoor''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amoor''))))) of + else if (((((string_startswith stringappend_16700 (''amoor''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amoor''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amoor''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amoor''))))) of s0 => Some (AMOOR, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16160 (''amomin''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amomin''))))) of + else if (((((string_startswith stringappend_16700 (''amomin''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amomin''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amomin''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amomin''))))) of s0 => Some (AMOMIN, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16160 (''amomax''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amomax''))))) of + else if (((((string_startswith stringappend_16700 (''amomax''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amomax''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amomax''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amomax''))))) of s0 => Some (AMOMAX, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16160 (''amominu''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amominu''))))) of + else if (((((string_startswith stringappend_16700 (''amominu''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amominu''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amominu''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amominu''))))) of s0 => Some (AMOMINU, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16160 (''amomaxu''))) \<and> ( - (case ((string_drop stringappend_16160 ((string_length (''amomaxu''))))) of + else if (((((string_startswith stringappend_16700 (''amomaxu''))) \<and> ( + (case ((string_drop stringappend_16700 ((string_length (''amomaxu''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16160 ((string_length (''amomaxu''))))) of + (case ((string_drop stringappend_16700 ((string_length (''amomaxu''))))) of s0 => Some (AMOMAXU, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -12301,9 +12335,9 @@ fun encdec_csrop_forwards :: " csrop \<Rightarrow>(2)Word.word " where definition encdec_csrop_backwards :: "(2)Word.word \<Rightarrow> csrop " where " encdec_csrop_backwards arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then CSRRW - else if (((p00 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then CSRRS + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then CSRRW + else if (((b__0 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then CSRRS else CSRRC))" @@ -12319,10 +12353,10 @@ fun encdec_csrop_forwards_matches :: " csrop \<Rightarrow> bool " where definition encdec_csrop_backwards_matches :: "(2)Word.word \<Rightarrow> bool " where " encdec_csrop_backwards_matches arg0 = ( - (let p00 = arg0 in - if (((p00 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then True - else if (((p00 = (vec_of_bits [B1,B1] :: 2 Word.word)))) then True + (let b__0 = arg0 in + if (((b__0 = (vec_of_bits [B0,B1] :: 2 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B0] :: 2 Word.word)))) then True + else if (((b__0 = (vec_of_bits [B1,B1] :: 2 Word.word)))) then True else False))" @@ -12361,7 +12395,7 @@ definition readCSR :: "(12)Word.word \<Rightarrow>((register_value),((64)Word.w else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (read_reg mscratch_ref :: ( 64 Word.word) M) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then - (read_reg mepc_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__12 :: 64 Word.word) . + (read_reg mepc_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__12 :: xlenbits) . (pc_alignment_mask () :: ( 64 Word.word) M) \<bind> (\<lambda> (w__13 :: 64 Word.word) . return ((and_vec w__12 w__13 :: 64 Word.word)))) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then @@ -12397,7 +12431,7 @@ definition readCSR :: "(12)Word.word \<Rightarrow>((register_value),((64)Word.w else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (read_reg sscratch_ref :: ( 64 Word.word) M) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then - (read_reg sepc_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__27 :: 64 Word.word) . + (read_reg sepc_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__27 :: xlenbits) . (pc_alignment_mask () :: ( 64 Word.word) M) \<bind> (\<lambda> (w__28 :: 64 Word.word) . return ((and_vec w__27 w__28 :: 64 Word.word)))) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then @@ -12418,7 +12452,7 @@ definition readCSR :: "(12)Word.word \<Rightarrow>((register_value),((64)Word.w else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (read_reg minstret_ref :: ( 64 Word.word) M) else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then - (read_reg tselect_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__37 :: 64 Word.word) . + (read_reg tselect_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__37 :: xlenbits) . return ((not_vec w__37 :: 64 Word.word))) else (let (_ :: unit) = (print_bits (''unhandled read to CSR '') csr) in @@ -12428,8 +12462,8 @@ definition readCSR :: "(12)Word.word \<Rightarrow>((register_value),((64)Word.w :: 64 Word.word))) \<bind> (\<lambda> (res :: xlenbits) . (let (_ :: unit) = (print_endline - (((op@) (''CSR '') - (((op@) ((csr_name csr)) (((op@) ('' -> '') ((string_of_bits res))))))))) in + (((@) (''CSR '') + (((@) ((csr_name csr)) (((@) ('' -> '') ((string_of_bits res))))))))) in return res))))" @@ -12475,18 +12509,18 @@ definition writeCSR :: "(12)Word.word \<Rightarrow>(64)Word.word \<Rightarrow>( return (Some ((EXTZ (( 64 :: int)::ii) ((get_Counteren_bits w__14 :: 32 Word.word)) :: 64 Word.word))))) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (write_reg mscratch_ref value1 \<then> - (read_reg mscratch_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__15 :: 64 Word.word) . return (Some w__15)) + (read_reg mscratch_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__15 :: xlenbits) . return (Some w__15)) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then - (legalize_xepc value1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__16 :: xlenbits) . + (legalize_xepc value1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__16 :: 64 Word.word) . (write_reg mepc_ref w__16 \<then> - (read_reg mepc_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__17 :: 64 Word.word) . return (Some w__17))) + (read_reg mepc_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__17 :: xlenbits) . return (Some w__17))) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (set_Mcause_bits mcause_ref value1 \<then> read_reg mcause_ref) \<bind> (\<lambda> (w__18 :: Mcause) . return (Some ((get_Mcause_bits w__18 :: 64 Word.word)))) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (write_reg mtval_ref value1 \<then> - (read_reg mtval_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__19 :: 64 Word.word) . return (Some w__19)) + (read_reg mtval_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__19 :: xlenbits) . return (Some w__19)) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then read_reg mip_ref \<bind> (\<lambda> (w__20 :: Minterrupts) . (write_reg mip_ref ((legalize_mip w__20 value1)) \<then> @@ -12494,10 +12528,10 @@ definition writeCSR :: "(12)Word.word \<Rightarrow>(64)Word.word \<Rightarrow>( return (Some ((get_Minterrupts_bits w__21 :: 64 Word.word))))) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (write_reg pmpcfg0_ref value1 \<then> - (read_reg pmpcfg0_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__22 :: 64 Word.word) . return (Some w__22)) + (read_reg pmpcfg0_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__22 :: xlenbits) . return (Some w__22)) else if (((b__0 = (vec_of_bits [B0,B0,B1,B1,B1,B0,B1,B1,B0,B0,B0,B0] :: 12 Word.word)))) then (write_reg pmpaddr0_ref value1 \<then> - (read_reg pmpaddr0_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__23 :: 64 Word.word) . return (Some w__23)) + (read_reg pmpaddr0_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__23 :: xlenbits) . return (Some w__23)) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then read_reg mstatus_ref \<bind> (\<lambda> (w__24 :: Mstatus) . (write_reg mstatus_ref ((legalize_sstatus w__24 value1)) \<then> @@ -12530,18 +12564,18 @@ definition writeCSR :: "(12)Word.word \<Rightarrow>(64)Word.word \<Rightarrow>( return (Some ((EXTZ (( 64 :: int)::ii) ((get_Counteren_bits w__35 :: 32 Word.word)) :: 64 Word.word))))) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (write_reg sscratch_ref value1 \<then> - (read_reg sscratch_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__36 :: 64 Word.word) . return (Some w__36)) + (read_reg sscratch_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__36 :: xlenbits) . return (Some w__36)) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B1] :: 12 Word.word)))) then - (legalize_xepc value1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__37 :: xlenbits) . + (legalize_xepc value1 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__37 :: 64 Word.word) . (write_reg sepc_ref w__37 \<then> - (read_reg sepc_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__38 :: 64 Word.word) . return (Some w__38))) + (read_reg sepc_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__38 :: xlenbits) . return (Some w__38))) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then (set_Mcause_bits scause_ref value1 \<then> read_reg scause_ref) \<bind> (\<lambda> (w__39 :: Mcause) . return (Some ((get_Mcause_bits w__39 :: 64 Word.word)))) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B1,B1] :: 12 Word.word)))) then (write_reg stval_ref value1 \<then> - (read_reg stval_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__40 :: 64 Word.word) . return (Some w__40)) + (read_reg stval_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__40 :: xlenbits) . return (Some w__40)) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B0,B1,B0,B0,B0,B1,B0,B0] :: 12 Word.word)))) then read_reg mip_ref \<bind> (\<lambda> (w__41 :: Minterrupts) . read_reg mideleg_ref \<bind> (\<lambda> (w__42 :: Minterrupts) . @@ -12550,28 +12584,28 @@ definition writeCSR :: "(12)Word.word \<Rightarrow>(64)Word.word \<Rightarrow>( return (Some ((get_Minterrupts_bits w__43 :: 64 Word.word)))))) else if (((b__0 = (vec_of_bits [B0,B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then cur_Architecture () \<bind> (\<lambda> (w__44 :: Architecture) . - (read_reg satp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__45 :: 64 Word.word) . + (read_reg satp_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__45 :: xlenbits) . (write_reg satp_ref ((legalize_satp w__44 w__45 value1 :: 64 Word.word)) \<then> - (read_reg satp_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__46 :: 64 Word.word) . return (Some w__46)))) + (read_reg satp_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__46 :: xlenbits) . return (Some w__46)))) else if (((b__0 = (vec_of_bits [B0,B1,B1,B1,B1,B0,B1,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (write_reg tselect_ref value1 \<then> - (read_reg tselect_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__47 :: 64 Word.word) . return (Some w__47)) + (read_reg tselect_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__47 :: xlenbits) . return (Some w__47)) else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word)))) then (write_reg mcycle_ref value1 \<then> - (read_reg mcycle_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__48 :: 64 Word.word) . return (Some w__48)) + (read_reg mcycle_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__48 :: xlenbits) . return (Some w__48)) else if (((b__0 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0] :: 12 Word.word)))) then ((write_reg minstret_ref value1 \<then> write_reg minstret_written_ref True) \<then> - (read_reg minstret_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__49 :: 64 Word.word) . return (Some w__49)) + (read_reg minstret_ref :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__49 :: xlenbits) . return (Some w__49)) else return None) \<bind> (\<lambda> (res :: xlenbits option) . return ((case res of Some (v) => print_endline - (((op@) (''CSR '') - (((op@) ((csr_name csr)) - (((op@) ('' <- '') - (((op@) ((string_of_bits v)) - (((op@) ('' (input: '') (((op@) ((string_of_bits value1)) ('')''))))))))))))) + (((@) (''CSR '') + (((@) ((csr_name csr)) + (((@) ('' <- '') + (((@) ((string_of_bits v)) + (((@) ('' (input: '') (((@) ((string_of_bits value1)) ('')''))))))))))))) | None => print_bits (''unhandled write to CSR '') csr )))))" @@ -12609,19 +12643,19 @@ definition maybe_i_backwards_matches :: " string \<Rightarrow> bool " where definition maybe_i_matches_prefix :: " string \<Rightarrow>(bool*int)option " where " maybe_i_matches_prefix arg0 = ( - (let stringappend_16140 = arg0 in - if (((((string_startswith stringappend_16140 (''i''))) \<and> ( - (case ((string_drop stringappend_16140 ((string_length (''i''))))) of + (let stringappend_16680 = arg0 in + if (((((string_startswith stringappend_16680 (''i''))) \<and> ( + (case ((string_drop stringappend_16680 ((string_length (''i''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16140 ((string_length (''i''))))) of + (case ((string_drop stringappend_16680 ((string_length (''i''))))) of s0 => Some (True, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16140 (''''))) \<and> ( - (case ((string_drop stringappend_16140 ((string_length (''''))))) of + else if (((((string_startswith stringappend_16680 (''''))) \<and> ( + (case ((string_drop stringappend_16680 ((string_length (''''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16140 ((string_length (''''))))) of + (case ((string_drop stringappend_16680 ((string_length (''''))))) of s0 => Some (False, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -12667,26 +12701,26 @@ definition csr_mnemonic_backwards_matches :: " string \<Rightarrow> bool " whe definition csr_mnemonic_matches_prefix :: " string \<Rightarrow>(csrop*int)option " where " csr_mnemonic_matches_prefix arg0 = ( - (let stringappend_16110 = arg0 in - if (((((string_startswith stringappend_16110 (''csrrw''))) \<and> ( - (case ((string_drop stringappend_16110 ((string_length (''csrrw''))))) of + (let stringappend_16650 = arg0 in + if (((((string_startswith stringappend_16650 (''csrrw''))) \<and> ( + (case ((string_drop stringappend_16650 ((string_length (''csrrw''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16110 ((string_length (''csrrw''))))) of + (case ((string_drop stringappend_16650 ((string_length (''csrrw''))))) of s0 => Some (CSRRW, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16110 (''csrrs''))) \<and> ( - (case ((string_drop stringappend_16110 ((string_length (''csrrs''))))) of + else if (((((string_startswith stringappend_16650 (''csrrs''))) \<and> ( + (case ((string_drop stringappend_16650 ((string_length (''csrrs''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16110 ((string_length (''csrrs''))))) of + (case ((string_drop stringappend_16650 ((string_length (''csrrs''))))) of s0 => Some (CSRRS, ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_16110 (''csrrc''))) \<and> ( - (case ((string_drop stringappend_16110 ((string_length (''csrrc''))))) of + else if (((((string_startswith stringappend_16650 (''csrrc''))) \<and> ( + (case ((string_drop stringappend_16650 ((string_length (''csrrc''))))) of s0 => True ))))) then - (case ((string_drop stringappend_16110 ((string_length (''csrrc''))))) of + (case ((string_drop stringappend_16650 ((string_length (''csrrc''))))) of s0 => Some (CSRRC, ((string_length arg0)) - ((string_length s0))) ) else None))" @@ -12701,71 +12735,71 @@ definition decodeCompressed :: "(16)Word.word \<Rightarrow>(ast)option " where Some (NOP () ) else None)) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word))))))) then - (let (nz54 :: 2 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 11 :: int)::ii) :: 2 Word.word)) in + (let (rd :: cregbits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let (nz96 :: 4 bits) = ((subrange_vec_dec v__2 (( 10 :: int)::ii) (( 7 :: int)::ii) :: 4 Word.word)) in - (let (nz2 :: 1 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in + (let (nz54 :: 2 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 11 :: int)::ii) :: 2 Word.word)) in (let (nz3 :: 1 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in - (let (rd :: cregbits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in + (let (nz2 :: 1 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let nzimm = ((concat_vec nz96 ((concat_vec nz54 ((concat_vec nz3 nz2 :: 2 Word.word)) :: 4 Word.word)) :: 8 Word.word)) in if (((nzimm = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word)))) then None else Some (C_ADDI4SPN (rd,nzimm)))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word))))))) then + (let (ui6 :: 1 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in - (let (rs1 :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in - (let (ui6 :: 1 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in + (let (rs1 :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (rd :: cregbits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let uimm = ((concat_vec ui6 ((concat_vec ui53 ui2 :: 4 Word.word)) :: 5 Word.word)) in Some (C_LW (uimm,rs1,rd)))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word))))))) then + (let (ui76 :: 2 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (rs1 :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in - (let (ui76 :: 2 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in (let (rd :: cregbits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let uimm = ((concat_vec ui76 ui53 :: 5 Word.word)) in Some (C_LD (uimm,rs1,rd))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word))))))) then + (let (ui6 :: 1 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (ui53 :: 3 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in - (let (rs1 :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (ui2 :: 1 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in - (let (ui6 :: 1 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (rs2 :: cregbits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in + (let (rs1 :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let uimm = ((concat_vec ui6 ((concat_vec ui53 ui2 :: 4 Word.word)) :: 5 Word.word)) in Some (C_SW (uimm,rs1,rs2)))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word))))))) then - (let (ui53 :: 3 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in - (let (rs1 :: 3 bits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let (ui76 :: 2 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in + (let (ui53 :: 3 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: 3 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in + (let (rs1 :: 3 bits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in (let uimm = ((concat_vec ui76 ui53 :: 5 Word.word)) in Some (C_SD (uimm,rs1,rs2))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then - (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (rsd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzi40 :: 5 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let nzi = ((concat_vec nzi5 nzi40 :: 6 Word.word)) in if ((((((nzi = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<or> (((((regbits_to_regno rsd)) = ((regbits_to_regno zreg)))))))) then None else Some (C_ADDI (nzi,rsd)))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then - (let (imm5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (rsd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm40 :: 5 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in Some (C_ADDIW ((concat_vec imm5 imm40 :: 6 Word.word),rsd))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then - (let (imm5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (rd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm40 :: 5 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in if (((((regbits_to_regno rd)) = ((regbits_to_regno zreg))))) then None else Some (C_LI ((concat_vec imm5 imm40 :: 6 Word.word),rd))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) \<and> ((((((((regbits_to_regno ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)))) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word)))))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))))))))) then (let (nzi9 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (nzi4 :: 1 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in - (let (nzi6 :: 1 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi87 :: 2 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in + (let (nzi6 :: 1 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 5 :: int)::ii) :: 1 Word.word)) in (let (nzi5 :: 1 bits) = ((subrange_vec_dec v__2 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in + (let (nzi4 :: 1 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in (let nzimm = ((concat_vec nzi9 ((concat_vec nzi87 ((concat_vec nzi6 ((concat_vec nzi5 nzi4 :: 2 Word.word)) :: 3 Word.word)) @@ -12774,29 +12808,29 @@ definition decodeCompressed :: "(16)Word.word \<Rightarrow>(ast)option " where if (((nzimm = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) then None else Some (C_ADDI16SP nzimm))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then - (let (imm17 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (rd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm17 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (imm1612 :: 5 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in if ((((((((regbits_to_regno rd)) = ((regbits_to_regno zreg))))) \<or> (((((regbits_to_regno rd)) = ((regbits_to_regno sp)))))))) then None else Some (C_LUI ((concat_vec imm17 imm1612 :: 6 Word.word),rd))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) \<and> ((((((((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))))))))) then - (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (rsd :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in + (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let (shamt :: 6 bits) = ((concat_vec nzui5 nzui40 :: 6 Word.word)) in if (((shamt = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) then None else Some (C_SRLI (shamt,rsd)))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) \<and> ((((((((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))))))))) then - (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (rsd :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in + (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let (shamt :: 6 bits) = ((concat_vec nzui5 nzui40 :: 6 Word.word)) in if (((shamt = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) then None else Some (C_SRAI (shamt,rsd)))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) \<and> ((((((((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))))))))) then - (let (i5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (rsd :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in + (let (i5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (i40 :: 5 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in Some (C_ANDI ((concat_vec i5 i40 :: 6 Word.word),rsd))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 10 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B1,B0,B0,B0,B1,B1] :: 6 Word.word)))) \<and> ((((((((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))))))))) then @@ -12824,14 +12858,14 @@ definition decodeCompressed :: "(16)Word.word \<Rightarrow>(ast)option " where (let (rs2 :: cregbits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in Some (C_ADDW (rsd,rs2)))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then - (let (i11 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (i4 :: 1 bits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in (let (i98 :: 2 bits) = ((subrange_vec_dec v__2 (( 10 :: int)::ii) (( 9 :: int)::ii) :: 2 Word.word)) in - (let (i10 :: 1 bits) = ((subrange_vec_dec v__2 (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word)) in - (let (i6 :: 1 bits) = ((subrange_vec_dec v__2 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (i7 :: 1 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 6 :: int)::ii) :: 1 Word.word)) in - (let (i31 :: 3 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) in + (let (i6 :: 1 bits) = ((subrange_vec_dec v__2 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__2 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in + (let (i4 :: 1 bits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in + (let (i31 :: 3 bits) = ((subrange_vec_dec v__2 (( 5 :: int)::ii) (( 3 :: int)::ii) :: 3 Word.word)) in + (let (i11 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (let (i10 :: 1 bits) = ((subrange_vec_dec v__2 (( 8 :: int)::ii) (( 8 :: int)::ii) :: 1 Word.word)) in Some (C_J ((concat_vec i11 ((concat_vec i10 ((concat_vec i98 @@ -12844,62 +12878,62 @@ definition decodeCompressed :: "(16)Word.word \<Rightarrow>(ast)option " where :: 10 Word.word)) :: 11 Word.word))))))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then - (let (i8 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (i43 :: 2 bits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) in (let (rs :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in + (let (i8 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (i76 :: 2 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in - (let (i21 :: 2 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__2 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in + (let (i43 :: 2 bits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) in + (let (i21 :: 2 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in Some (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__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then - (let (i8 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (i43 :: 2 bits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) in (let (rs :: cregbits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in + (let (i8 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (i76 :: 2 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in - (let (i21 :: 2 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in (let (i5 :: 1 bits) = ((subrange_vec_dec v__2 (( 2 :: int)::ii) (( 2 :: int)::ii) :: 1 Word.word)) in + (let (i43 :: 2 bits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 10 :: int)::ii) :: 2 Word.word)) in + (let (i21 :: 2 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 3 :: int)::ii) :: 2 Word.word)) in Some (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 ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then - (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (rsd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (nzui5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let (nzui40 :: 5 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let (shamt :: 6 bits) = ((concat_vec nzui5 nzui40 :: 6 Word.word)) in if ((((((shamt = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<or> (((((regbits_to_regno rsd)) = ((regbits_to_regno zreg)))))))) then None else Some (C_SLLI (shamt,rsd)))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then + (let (ui76 :: 2 bits) = ((subrange_vec_dec v__2 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in (let (ui5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (rd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (ui42 :: 3 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 4 :: int)::ii) :: 3 Word.word)) in - (let (ui76 :: 2 bits) = ((subrange_vec_dec v__2 (( 3 :: int)::ii) (( 2 :: int)::ii) :: 2 Word.word)) in + (let (rd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (uimm :: 6 bits) = ((concat_vec ui76 ((concat_vec ui5 ui42 :: 4 Word.word)) :: 6 Word.word)) in if (((((regbits_to_regno rd)) = ((regbits_to_regno zreg))))) then None else Some (C_LWSP (uimm,rd))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then + (let (ui86 :: 3 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in (let (ui5 :: 1 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (rd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (ui43 :: 2 bits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in - (let (ui86 :: 3 bits) = ((subrange_vec_dec v__2 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in + (let (rd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (uimm :: 6 bits) = ((concat_vec ui86 ((concat_vec ui5 ui43 :: 3 Word.word)) :: 6 Word.word)) in if (((((regbits_to_regno rd)) = ((regbits_to_regno zreg))))) then None else Some (C_LDSP (uimm,rd))))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then - (let (ui52 :: 4 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 9 :: int)::ii) :: 4 Word.word)) in (let (ui76 :: 2 bits) = ((subrange_vec_dec v__2 (( 8 :: int)::ii) (( 7 :: int)::ii) :: 2 Word.word)) in + (let (ui52 :: 4 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 9 :: int)::ii) :: 4 Word.word)) in (let (rs2 :: regbits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let (uimm :: 6 bits) = ((concat_vec ui76 ui52 :: 6 Word.word)) in Some (C_SWSP (uimm,rs2)))))) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 13 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then - (let (ui53 :: 3 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (ui86 :: 3 bits) = ((subrange_vec_dec v__2 (( 9 :: int)::ii) (( 7 :: int)::ii) :: 3 Word.word)) in + (let (ui53 :: 3 bits) = ((subrange_vec_dec v__2 (( 12 :: int)::ii) (( 10 :: int)::ii) :: 3 Word.word)) in (let (rs2 :: regbits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in (let (uimm :: 6 bits) = ((concat_vec ui86 ui53 :: 6 Word.word)) in Some (C_SDSP (uimm,rs2)))))) @@ -12912,8 +12946,8 @@ definition decodeCompressed :: "(16)Word.word \<Rightarrow>(ast)option " where if (((((regbits_to_regno rs1)) = ((regbits_to_regno zreg))))) then None else Some (C_JALR rs1)) else if ((((((((subrange_vec_dec v__2 (( 15 :: int)::ii) (( 12 :: int)::ii) :: 4 Word.word)) = (vec_of_bits [B1,B0,B0,B0] :: 4 Word.word)))) \<and> (((((subrange_vec_dec v__2 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then - (let (rd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in (let (rs2 :: regbits) = ((subrange_vec_dec v__2 (( 6 :: int)::ii) (( 2 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regbits) = ((subrange_vec_dec v__2 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in if ((((((((regbits_to_regno rs2)) = ((regbits_to_regno zreg))))) \<or> (((((regbits_to_regno rd)) = ((regbits_to_regno zreg)))))))) then None else Some (C_MV (rd,rs2)))) @@ -12931,7 +12965,7 @@ definition decodeCompressed :: "(16)Word.word \<Rightarrow>(ast)option " where (*val execute_WFI : unit -> M bool*) fun execute_WFI :: " unit \<Rightarrow>((register_value),(bool),(exception))monad " where - " execute_WFI g__26 = ( + " execute_WFI _ = ( read_reg cur_privilege_ref \<bind> (\<lambda> (w__0 :: Privilege) . (case w__0 of Machine => return True @@ -12956,7 +12990,7 @@ fun execute_UTYPE :: "(20)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> uo (case op1 of RISCV_LUI => return off | RISCV_AUIPC => - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . return ((add_vec w__0 off :: 64 Word.word))) ) \<bind> (\<lambda> (ret :: xlenbits) . wX ((regbits_to_regno rd)) ret \<then> return True)))" @@ -12965,7 +12999,7 @@ fun execute_UTYPE :: "(20)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> uo (*val execute_THREAD_START : unit -> bool*) fun execute_THREAD_START :: " unit \<Rightarrow> bool " where - " execute_THREAD_START g__29 = ( True )" + " execute_THREAD_START _ = ( True )" (*val execute_STORECON : bool -> bool -> mword ty5 -> mword ty5 -> word_width -> mword ty5 -> M bool*) @@ -13009,13 +13043,18 @@ fun execute_STORECON :: " bool \<Rightarrow> bool \<Rightarrow>(5)Word.word \< ((subrange_vec_dec rs2_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) aq rl True | DOUBLE => mem_write_value addr (( 8 :: int)::ii) rs2_val aq rl True | _ => internal_error (''STORECON expected word or double'') - ) \<bind> (\<lambda> (res :: unit MemoryOpResult) . + ) \<bind> (\<lambda> (res :: bool MemoryOpResult) . (case res of - MemValue (_) => + MemValue (True) => wX ((regbits_to_regno rd)) ((EXTZ (( 64 :: int)::ii) (vec_of_bits [B0] :: 1 Word.word) :: 64 Word.word)) \<then> ((let (_ :: unit) = (cancel_reservation () ) in return True)) + | MemValue (False) => + wX ((regbits_to_regno rd)) + ((EXTZ (( 64 :: int)::ii) (vec_of_bits [B1] :: 1 Word.word) :: 64 Word.word)) \<then> + ((let (_ :: unit) = (cancel_reservation () ) in + return True)) | MemException (e) => handle_mem_exception addr e \<then> return False ))) )) @@ -13056,9 +13095,10 @@ fun execute_STORE :: "(12)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5) mem_write_value addr (( 4 :: int)::ii) ((subrange_vec_dec rs2_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) aq rl False | DOUBLE => mem_write_value addr (( 8 :: int)::ii) rs2_val aq rl False - ) \<bind> (\<lambda> (res :: unit MemoryOpResult) . + ) \<bind> (\<lambda> (res :: bool MemoryOpResult) . (case res of - MemValue (_) => return True + MemValue (True) => return True + | MemValue (False) => internal_error (''store got false from mem_write_value'') | MemException (e) => handle_mem_exception addr e \<then> return False ))) )) @@ -13068,13 +13108,13 @@ fun execute_STORE :: "(12)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5) (*val execute_STOP_FETCHING : unit -> bool*) fun execute_STOP_FETCHING :: " unit \<Rightarrow> bool " where - " execute_STOP_FETCHING g__28 = ( True )" + " execute_STOP_FETCHING _ = ( True )" (*val execute_SRET : unit -> M bool*) fun execute_SRET :: " unit \<Rightarrow>((register_value),(bool),(exception))monad " where - " execute_SRET g__24 = ( + " execute_SRET _ = ( read_reg cur_privilege_ref \<bind> (\<lambda> (w__0 :: Privilege) . (case w__0 of User => handle_illegal () @@ -13084,13 +13124,13 @@ fun execute_SRET :: " unit \<Rightarrow>((register_value),(bool),(exception))m handle_illegal () else read_reg cur_privilege_ref \<bind> (\<lambda> (w__2 :: Privilege) . - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . - (handle_exception w__2 (CTL_SRET () ) w__3 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: xlenbits) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: xlenbits) . + (handle_exception w__2 (CTL_SRET () ) w__3 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__4 :: 64 Word.word) . write_reg nextPC_ref w__4)))) | Machine => read_reg cur_privilege_ref \<bind> (\<lambda> (w__5 :: Privilege) . - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: 64 Word.word) . - (handle_exception w__5 (CTL_SRET () ) w__6 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: xlenbits) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: xlenbits) . + (handle_exception w__5 (CTL_SRET () ) w__6 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: 64 Word.word) . write_reg nextPC_ref w__7))) ) \<then> return False))" @@ -13111,6 +13151,32 @@ fun execute_SHIFTW :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5) wX ((regbits_to_regno rd)) ((EXTS (( 64 :: int)::ii) result :: 64 Word.word)) \<then> return True))))" +(*val execute_SHIFTIWOP : mword ty5 -> mword ty5 -> mword ty5 -> sopw -> M bool*) + +fun execute_SHIFTIWOP :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> sopw \<Rightarrow>((register_value),(bool),(exception))monad " where + " execute_SHIFTIWOP shamt rs1 rd op1 = ( + (rX ((regbits_to_regno rs1)) :: ( 64 Word.word) M) \<bind> (\<lambda> rs1_val . + (let (result :: xlenbits) = + ((case op1 of + RISCV_SLLIW => + (EXTS (( 64 :: int)::ii) + ((shift_bits_left ((subrange_vec_dec rs1_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) shamt + :: 32 Word.word)) + :: 64 Word.word) + | RISCV_SRLIW => + (EXTS (( 64 :: int)::ii) + ((shift_bits_right ((subrange_vec_dec rs1_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) shamt + :: 32 Word.word)) + :: 64 Word.word) + | RISCV_SRAIW => + (EXTS (( 64 :: int)::ii) + ((shift_right_arith32 ((subrange_vec_dec rs1_val (( 31 :: int)::ii) (( 0 :: int)::ii) :: 32 Word.word)) shamt + :: 32 Word.word)) + :: 64 Word.word) + )) in + wX ((regbits_to_regno rd)) result \<then> return True)))" + + (*val execute_SHIFTIOP : mword ty6 -> mword ty5 -> mword ty5 -> sop -> M bool*) fun execute_SHIFTIOP :: "(6)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> sop \<Rightarrow>((register_value),(bool),(exception))monad " where @@ -13136,9 +13202,9 @@ fun execute_SFENCE_VMA :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow else read_reg mstatus_ref \<bind> (\<lambda> (w__1 :: Mstatus) . read_reg mstatus_ref \<bind> (\<lambda> (w__2 :: Mstatus) . - (let p__20 = + (let p__16 = (architecture ((get_Mstatus_SXL w__1 :: 2 Word.word)), (get_Mstatus_TVM w__2 :: 1 Word.word)) in - (case p__20 of + (case p__16 of (Some (RV64), v_0) => if (((v_0 = ((bool_to_bits True :: 1 Word.word))))) then handle_illegal () \<then> return False else if (((v_0 = ((bool_to_bits False :: 1 Word.word))))) then @@ -13155,9 +13221,9 @@ fun execute_SFENCE_VMA :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow flushTLB asid addr \<then> return True)) else (case (Some RV64, v_0) of - (g__18, g__19) => internal_error (''unimplemented sfence architecture'') + (g__14, g__15) => internal_error (''unimplemented sfence architecture'') ) - | (g__18, g__19) => internal_error (''unimplemented sfence architecture'') + | (g__14, g__15) => internal_error (''unimplemented sfence architecture'') ))))))" @@ -13220,7 +13286,7 @@ fun execute_RTYPE :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)W fun execute_RISCV_JALR :: "(12)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>((register_value),(bool),(exception))monad " where " execute_RISCV_JALR imm rs1 rd = ( - (read_reg nextPC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg nextPC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . (wX ((regbits_to_regno rd)) w__0 \<then> (rX ((regbits_to_regno rs1)) :: ( 64 Word.word) M)) \<bind> (\<lambda> (w__1 :: 64 Word.word) . (let (newPC :: xlenbits) = ((add_vec w__1 ((EXTS (( 64 :: int)::ii) imm :: 64 Word.word)) :: 64 Word.word)) in @@ -13237,7 +13303,7 @@ fun execute_RISCV_JALR :: "(12)Word.word \<Rightarrow>(5)Word.word \<Rightarro fun execute_RISCV_JAL :: "(21)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>((register_value),(bool),(exception))monad " where " execute_RISCV_JAL imm rd = ( (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (pc :: xlenbits) . - (read_reg nextPC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg nextPC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . wX ((regbits_to_regno rd)) w__0 \<then> ((let (offset :: xlenbits) = ((EXTS (( 64 :: int)::ii) imm :: 64 Word.word)) in write_reg nextPC_ref ((add_vec pc offset :: 64 Word.word)) \<then> return True)))))" @@ -13273,7 +13339,7 @@ fun execute_REM :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)Wor (*val execute_NOP : unit -> bool*) fun execute_NOP :: " unit \<Rightarrow> bool " where - " execute_NOP g__27 = ( True )" + " execute_NOP _ = ( True )" (*val execute_MULW : mword ty5 -> mword ty5 -> mword ty5 -> M bool*) @@ -13312,13 +13378,13 @@ fun execute_MUL :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)Wor (*val execute_MRET : unit -> M bool*) fun execute_MRET :: " unit \<Rightarrow>((register_value),(bool),(exception))monad " where - " execute_MRET g__23 = ( + " execute_MRET _ = ( read_reg cur_privilege_ref \<bind> (\<lambda> (w__0 :: Privilege) . (if (((((privLevel_to_bits w__0 :: 2 Word.word)) = ((privLevel_to_bits Machine :: 2 Word.word))))) then read_reg cur_privilege_ref \<bind> (\<lambda> (w__1 :: Privilege) . - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . - (handle_exception w__1 (CTL_MRET () ) w__2 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: xlenbits) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: xlenbits) . + (handle_exception w__1 (CTL_MRET () ) w__2 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . write_reg nextPC_ref w__3))) else handle_illegal () ) \<then> return False))" @@ -13339,7 +13405,7 @@ fun execute_LOADRES :: " bool \<Rightarrow> bool \<Rightarrow>(5)Word.word \<R | DOUBLE => (((subrange_vec_dec vaddr (( 2 :: int)::ii) (( 0 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)) )) in - if ((\<not> aligned)) then handle_mem_exception vaddr E_SAMO_Addr_Align \<then> return False + if ((\<not> aligned)) then handle_mem_exception vaddr E_Load_Addr_Align \<then> return False else translateAddr vaddr Read Data \<bind> (\<lambda> (w__0 :: TR_Result) . (case w__0 of @@ -13419,10 +13485,10 @@ fun execute_ILLEGAL :: "(32)Word.word \<Rightarrow>((register_value),(bool),(e " execute_ILLEGAL s = ( handle_illegal () \<then> return False )" -(*val execute_FENCEI : unit -> M bool*) +(*val execute_FENCEI : unit -> bool*) -fun execute_FENCEI :: " unit \<Rightarrow>((register_value),(bool),(exception))monad " where - " execute_FENCEI g__21 = ( MEM_fence_i () \<then> return True )" +fun execute_FENCEI :: " unit \<Rightarrow> bool " where + " execute_FENCEI _ = ( True )" (*val execute_FENCE : mword ty4 -> mword ty4 -> M bool*) @@ -13462,7 +13528,7 @@ fun execute_FENCE :: "(4)Word.word \<Rightarrow>(4)Word.word \<Rightarrow>((re (*val execute_ECALL : unit -> M bool*) fun execute_ECALL :: " unit \<Rightarrow>((register_value),(bool),(exception))monad " where - " execute_ECALL g__22 = ( + " execute_ECALL _ = ( read_reg cur_privilege_ref \<bind> (\<lambda> (w__0 :: Privilege) . (let (t :: sync_exception) = ((| sync_exception_trap = @@ -13473,16 +13539,16 @@ fun execute_ECALL :: " unit \<Rightarrow>((register_value),(bool),(exception)) )), sync_exception_excinfo = None |)) in read_reg cur_privilege_ref \<bind> (\<lambda> (w__1 :: Privilege) . - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: 64 Word.word) . - (handle_exception w__1 (CTL_TRAP t) w__2 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: xlenbits) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__2 :: xlenbits) . + (handle_exception w__1 (CTL_TRAP t) w__2 :: ( 64 Word.word) M) \<bind> (\<lambda> (w__3 :: 64 Word.word) . write_reg nextPC_ref w__3 \<then> return False))))))" (*val execute_EBREAK : unit -> M bool*) fun execute_EBREAK :: " unit \<Rightarrow>((register_value),(bool),(exception))monad " where - " execute_EBREAK g__25 = ( - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + " execute_EBREAK _ = ( + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . handle_mem_exception w__0 E_Breakpoint \<then> return False))" @@ -13521,7 +13587,7 @@ fun execute_DIV :: "(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)Wor (*val execute_C_ILLEGAL : unit -> M bool*) fun execute_C_ILLEGAL :: " unit \<Rightarrow>((register_value),(bool),(exception))monad " where - " execute_C_ILLEGAL g__30 = ( handle_illegal () \<then> return False )" + " execute_C_ILLEGAL _ = ( handle_illegal () \<then> return False )" (*val execute_C_ADDIW : mword ty6 -> mword ty5 -> M bool*) @@ -13579,7 +13645,7 @@ fun execute_BTYPE :: "(13)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5) | RISCV_BGEU => zopz0zKzJ_u rs1_val rs2_val )) in (if taken then - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . write_reg nextPC_ref ((add_vec w__0 ((EXTS (( 64 :: int)::ii) imm :: 64 Word.word)) :: 64 Word.word))) else return () ) \<then> return True))))" @@ -13635,9 +13701,10 @@ fun execute_AMO :: " amoop \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> (((aq \<and> rl))) rl True | DOUBLE => mem_write_value addr (( 8 :: int)::ii) result (((aq \<and> rl))) rl True | _ => internal_error (''AMO expected WORD or DOUBLE'') - ) \<bind> (\<lambda> (wval :: unit MemoryOpResult) . + ) \<bind> (\<lambda> (wval :: bool MemoryOpResult) . (case wval of - MemValue (_) => wX ((regbits_to_regno rd)) loaded \<then> return True + MemValue (True) => wX ((regbits_to_regno rd)) loaded \<then> return True + | MemValue (False) => internal_error (''AMO got false from mem_write_value'') | MemException (e) => handle_mem_exception addr e \<then> return False )))) )) @@ -13794,6 +13861,7 @@ function (sequential,domintros) execute :: " ast \<Rightarrow>((register_value |" execute (ADDIW (imm,rs1,rd)) = ( execute_ADDIW imm rs1 rd )" |" execute (SHIFTW (shamt,rs1,rd,op1)) = ( execute_SHIFTW shamt rs1 rd op1 )" |" execute (RTYPEW (rs2,rs1,rd,op1)) = ( execute_RTYPEW rs2 rs1 rd op1 )" +|" execute (SHIFTIWOP (shamt,rs1,rd,op1)) = ( execute_SHIFTIWOP shamt rs1 rd op1 )" |" execute (MUL (rs2,rs1,rd,high,signed1,signed2)) = ( execute_MUL rs2 rs1 rd high signed1 signed2 )" |" execute (DIV (rs2,rs1,rd,s)) = ( execute_DIV rs2 rs1 rd s )" |" execute (REM (rs2,rs1,rd,s)) = ( execute_REM rs2 rs1 rd s )" @@ -13801,23 +13869,23 @@ function (sequential,domintros) execute :: " ast \<Rightarrow>((register_value |" execute (DIVW (rs2,rs1,rd,s)) = ( execute_DIVW rs2 rs1 rd s )" |" execute (REMW (rs2,rs1,rd,s)) = ( execute_REMW rs2 rs1 rd s )" |" execute (FENCE (pred,succ)) = ( execute_FENCE pred succ )" -|" execute (FENCEI (g__21)) = ( execute_FENCEI g__21 )" -|" execute (ECALL (g__22)) = ( execute_ECALL g__22 )" -|" execute (MRET (g__23)) = ( execute_MRET g__23 )" -|" execute (SRET (g__24)) = ( execute_SRET g__24 )" -|" execute (EBREAK (g__25)) = ( execute_EBREAK g__25 )" -|" execute (WFI (g__26)) = ( execute_WFI g__26 )" +|" execute (FENCEI (arg0)) = ( return ((execute_FENCEI arg0)))" +|" execute (ECALL (arg0)) = ( execute_ECALL arg0 )" +|" execute (MRET (arg0)) = ( execute_MRET arg0 )" +|" execute (SRET (arg0)) = ( execute_SRET arg0 )" +|" execute (EBREAK (arg0)) = ( execute_EBREAK arg0 )" +|" execute (WFI (arg0)) = ( execute_WFI arg0 )" |" execute (SFENCE_VMA (rs1,rs2)) = ( execute_SFENCE_VMA rs1 rs2 )" |" execute (LOADRES (aq,rl,rs1,width,rd)) = ( execute_LOADRES aq rl rs1 width rd )" |" execute (STORECON (aq,rl,rs2,rs1,width,rd)) = ( execute_STORECON aq rl rs2 rs1 width rd )" |" execute (AMO (op1,aq,rl,rs2,rs1,width,rd)) = ( execute_AMO op1 aq rl rs2 rs1 width rd )" |" execute (CSR (csr,rs1,rd,is_imm,op1)) = ( execute_CSR csr rs1 rd is_imm op1 )" -|" execute (NOP (g__27)) = ( return ((execute_NOP g__27)))" +|" execute (NOP (arg0)) = ( return ((execute_NOP arg0)))" |" execute (C_ADDIW (imm,rsd)) = ( execute_C_ADDIW imm rsd )" -|" execute (STOP_FETCHING (g__28)) = ( return ((execute_STOP_FETCHING g__28)))" -|" execute (THREAD_START (g__29)) = ( return ((execute_THREAD_START g__29)))" +|" execute (STOP_FETCHING (arg0)) = ( return ((execute_STOP_FETCHING arg0)))" +|" execute (THREAD_START (arg0)) = ( return ((execute_THREAD_START arg0)))" |" execute (ILLEGAL (s)) = ( execute_ILLEGAL s )" -|" execute (C_ILLEGAL (g__30)) = ( execute_C_ILLEGAL g__30 )" +|" execute (C_ILLEGAL (arg0)) = ( execute_C_ILLEGAL arg0 )" by pat_completeness auto @@ -13920,6 +13988,13 @@ fun assembly_forwards :: " ast \<Rightarrow> string " where ((string_append ((reg_name_forwards rs1)) ((string_append ((sep_forwards () )) ((string_append ((reg_name_forwards rs2)) (''''))))))))))))))" +|" assembly_forwards (SHIFTIWOP (shamt,rs1,rd,op1)) = ( + string_append ((shiftiwop_mnemonic_forwards op1)) + ((string_append ((spc_forwards () )) + ((string_append ((reg_name_forwards rd)) + ((string_append ((sep_forwards () )) + ((string_append ((reg_name_forwards rs1)) + ((string_append ((string_of_bits shamt)) (''''))))))))))))" |" assembly_forwards (MUL (rs2,rs1,rd,high,signed1,signed2)) = ( string_append ((mul_mnemonic_forwards high signed1 signed2)) ((string_append ((spc_forwards () )) @@ -14049,24 +14124,24 @@ fun assembly_forwards :: " ast \<Rightarrow> string " where definition assembly_backwards :: " string \<Rightarrow> ast " where " assembly_backwards arg0 = ( - (let stringappend_10760 = arg0 in - if ((case ((utype_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_10770,stringappend_10780) => - (let stringappend_10790 = (string_drop stringappend_10760 stringappend_10780) in - if ((case ((spc_matches_prefix stringappend_10790)) of - Some (stringappend_10800,stringappend_10810) => - (let stringappend_10820 = (string_drop stringappend_10790 stringappend_10810) in - if ((case ((reg_name_matches_prefix stringappend_10820 :: (( 5 Word.word * ii))option)) of - Some (stringappend_10830,stringappend_10840) => - (let stringappend_10850 = (string_drop stringappend_10820 stringappend_10840) in - if ((case ((sep_matches_prefix stringappend_10850)) of - Some (stringappend_10860,stringappend_10870) => - (let stringappend_10880 = (string_drop stringappend_10850 stringappend_10870) in + (let stringappend_11120 = arg0 in + if ((case ((utype_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_11130,stringappend_11140) => + (let stringappend_11150 = (string_drop stringappend_11120 stringappend_11140) in + if ((case ((spc_matches_prefix stringappend_11150)) of + Some (stringappend_11160,stringappend_11170) => + (let stringappend_11180 = (string_drop stringappend_11150 stringappend_11170) in + if ((case ((reg_name_matches_prefix stringappend_11180 :: (( 5 Word.word * ii))option)) of + Some (stringappend_11190,stringappend_11200) => + (let stringappend_11210 = (string_drop stringappend_11180 stringappend_11200) in + if ((case ((sep_matches_prefix stringappend_11210)) of + Some (stringappend_11220,stringappend_11230) => + (let stringappend_11240 = (string_drop stringappend_11210 stringappend_11230) in if ((case ((hex_bits_20_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_10880 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11240 :: (( 20 Word.word * ii))option)) of - Some (stringappend_10890,stringappend_10900) => - if(((string_drop stringappend_10880 stringappend_10900)) = ('''')) then + Some (stringappend_11250,stringappend_11260) => + if(((string_drop stringappend_11240 stringappend_11260)) = ('''')) then True else False | None => False )) then @@ -14086,64 +14161,64 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where else False) | None => False )) then - (let (op1, stringappend_10780) = - ((case ((utype_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_10770,stringappend_10780) => - (stringappend_10770, stringappend_10780) + (let (op1, stringappend_11140) = + ((case ((utype_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_11130,stringappend_11140) => + (stringappend_11130, stringappend_11140) )) in - (let stringappend_10790 = (string_drop stringappend_10760 stringappend_10780) in + (let stringappend_11150 = (string_drop stringappend_11120 stringappend_11140) in (case - (case ((spc_matches_prefix stringappend_10790)) of - Some (stringappend_10800,stringappend_10810) => - (stringappend_10800, stringappend_10810) + (case ((spc_matches_prefix stringappend_11150)) of + Some (stringappend_11160,stringappend_11170) => + (stringappend_11160, stringappend_11170) ) of - (_, stringappend_10810) => - (let stringappend_10820 = (string_drop stringappend_10790 - stringappend_10810) in - (let (rd, stringappend_10840) = - ((case ((reg_name_matches_prefix stringappend_10820 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10830,stringappend_10840) => - (stringappend_10830, stringappend_10840) + (_, stringappend_11170) => + (let stringappend_11180 = (string_drop stringappend_11150 + stringappend_11170) in + (let (rd, stringappend_11200) = + ((case ((reg_name_matches_prefix stringappend_11180 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_11190,stringappend_11200) => + (stringappend_11190, stringappend_11200) )) in - (let stringappend_10850 = (string_drop stringappend_10820 - stringappend_10840) in + (let stringappend_11210 = (string_drop stringappend_11180 + stringappend_11200) in (case - (case ((sep_matches_prefix stringappend_10850)) of - Some (stringappend_10860,stringappend_10870) => - (stringappend_10860, stringappend_10870) + (case ((sep_matches_prefix stringappend_11210)) of + Some (stringappend_11220,stringappend_11230) => + (stringappend_11220, stringappend_11230) ) of - (_, stringappend_10870) => - (let stringappend_10880 = (string_drop stringappend_10850 - stringappend_10870) in - (let (imm, stringappend_10900) = + (_, stringappend_11230) => + (let stringappend_11240 = (string_drop stringappend_11210 + stringappend_11230) in + (let (imm, stringappend_11260) = ((case ((hex_bits_20_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_10880 :: (( 20 Word.word * ii)) option)) of - Some (stringappend_10890,stringappend_10900) => - (stringappend_10890, stringappend_10900) + stringappend_11240 :: (( 20 Word.word * ii)) option)) of + Some (stringappend_11250,stringappend_11260) => + (stringappend_11250, stringappend_11260) )) in - if(((string_drop stringappend_10880 stringappend_10900)) = ('''')) then + if(((string_drop stringappend_11240 stringappend_11260)) = ('''')) then (UTYPE (imm,rd,op1)) else undefined)) )))) ))) - else if (((((string_startswith stringappend_10760 (''jal''))) \<and> ((let stringappend_10920 = (string_drop stringappend_10760 ((string_length (''jal'')))) in - if ((case ((spc_matches_prefix stringappend_10920)) of - Some (stringappend_10930,stringappend_10940) => - (let stringappend_10950 = (string_drop stringappend_10920 stringappend_10940) in - if ((case ((reg_name_matches_prefix stringappend_10950 + else if (((((string_startswith stringappend_11120 (''jal''))) \<and> ((let stringappend_11280 = (string_drop stringappend_11120 ((string_length (''jal'')))) in + if ((case ((spc_matches_prefix stringappend_11280)) of + Some (stringappend_11290,stringappend_11300) => + (let stringappend_11310 = (string_drop stringappend_11280 stringappend_11300) in + if ((case ((reg_name_matches_prefix stringappend_11310 :: (( 5 Word.word * ii))option)) of - Some (stringappend_10960,stringappend_10970) => - (let stringappend_10980 = - (string_drop stringappend_10950 stringappend_10970) in - if ((case ((sep_matches_prefix stringappend_10980)) of - Some (stringappend_10990,stringappend_11000) => - (let stringappend_11010 = - (string_drop stringappend_10980 stringappend_11000) in + Some (stringappend_11320,stringappend_11330) => + (let stringappend_11340 = + (string_drop stringappend_11310 stringappend_11330) in + if ((case ((sep_matches_prefix stringappend_11340)) of + Some (stringappend_11350,stringappend_11360) => + (let stringappend_11370 = + (string_drop stringappend_11340 stringappend_11360) in if ((case ((hex_bits_21_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11010 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11370 :: (( 21 Word.word * ii))option)) of - Some (stringappend_11020,stringappend_11030) => - if(((string_drop stringappend_11010 stringappend_11030)) = ('''')) then + Some (stringappend_11380,stringappend_11390) => + if(((string_drop stringappend_11370 stringappend_11390)) = ('''')) then True else False | None => False )) then @@ -14161,68 +14236,68 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_10920 = (string_drop stringappend_10760 ((string_length (''jal'')))) in + (let stringappend_11280 = (string_drop stringappend_11120 ((string_length (''jal'')))) in (case - (case ((spc_matches_prefix stringappend_10920)) of - Some (stringappend_10930,stringappend_10940) => - (stringappend_10930, stringappend_10940) + (case ((spc_matches_prefix stringappend_11280)) of + Some (stringappend_11290,stringappend_11300) => + (stringappend_11290, stringappend_11300) ) of - (_, stringappend_10940) => - (let stringappend_10950 = (string_drop stringappend_10920 - stringappend_10940) in - (let (rd, stringappend_10970) = - ((case ((reg_name_matches_prefix stringappend_10950 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10960,stringappend_10970) => - (stringappend_10960, stringappend_10970) + (_, stringappend_11300) => + (let stringappend_11310 = (string_drop stringappend_11280 + stringappend_11300) in + (let (rd, stringappend_11330) = + ((case ((reg_name_matches_prefix stringappend_11310 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_11320,stringappend_11330) => + (stringappend_11320, stringappend_11330) )) in - (let stringappend_10980 = (string_drop stringappend_10950 - stringappend_10970) in + (let stringappend_11340 = (string_drop stringappend_11310 + stringappend_11330) in (case - (case ((sep_matches_prefix stringappend_10980)) of - Some (stringappend_10990,stringappend_11000) => - (stringappend_10990, stringappend_11000) + (case ((sep_matches_prefix stringappend_11340)) of + Some (stringappend_11350,stringappend_11360) => + (stringappend_11350, stringappend_11360) ) of - (_, stringappend_11000) => - (let stringappend_11010 = (string_drop stringappend_10980 - stringappend_11000) in - (let (imm, stringappend_11030) = + (_, stringappend_11360) => + (let stringappend_11370 = (string_drop stringappend_11340 + stringappend_11360) in + (let (imm, stringappend_11390) = ((case ((hex_bits_21_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_11010 :: (( 21 Word.word * ii)) option)) of - Some (stringappend_11020,stringappend_11030) => - (stringappend_11020, stringappend_11030) + stringappend_11370 :: (( 21 Word.word * ii)) option)) of + Some (stringappend_11380,stringappend_11390) => + (stringappend_11380, stringappend_11390) )) in - if(((string_drop stringappend_11010 stringappend_11030)) = ('''')) then + if(((string_drop stringappend_11370 stringappend_11390)) = ('''')) then (RISCV_JAL (imm,rd)) else undefined)) )))) )) - else if (((((string_startswith stringappend_10760 (''jalr''))) \<and> ((let stringappend_11050 = (string_drop stringappend_10760 ((string_length (''jalr'')))) in - if ((case ((spc_matches_prefix stringappend_11050)) of - Some (stringappend_11060,stringappend_11070) => - (let stringappend_11080 = (string_drop stringappend_11050 stringappend_11070) in - if ((case ((reg_name_matches_prefix stringappend_11080 + else if (((((string_startswith stringappend_11120 (''jalr''))) \<and> ((let stringappend_11410 = (string_drop stringappend_11120 ((string_length (''jalr'')))) in + if ((case ((spc_matches_prefix stringappend_11410)) of + Some (stringappend_11420,stringappend_11430) => + (let stringappend_11440 = (string_drop stringappend_11410 stringappend_11430) in + if ((case ((reg_name_matches_prefix stringappend_11440 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11090,stringappend_11100) => - (let stringappend_11110 = - (string_drop stringappend_11080 stringappend_11100) in - if ((case ((sep_matches_prefix stringappend_11110)) of - Some (stringappend_11120,stringappend_11130) => - (let stringappend_11140 = - (string_drop stringappend_11110 stringappend_11130) in - if ((case ((reg_name_matches_prefix stringappend_11140 + Some (stringappend_11450,stringappend_11460) => + (let stringappend_11470 = + (string_drop stringappend_11440 stringappend_11460) in + if ((case ((sep_matches_prefix stringappend_11470)) of + Some (stringappend_11480,stringappend_11490) => + (let stringappend_11500 = + (string_drop stringappend_11470 stringappend_11490) in + if ((case ((reg_name_matches_prefix stringappend_11500 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11150,stringappend_11160) => - (let stringappend_11170 = - (string_drop stringappend_11140 stringappend_11160) in - if ((case ((sep_matches_prefix stringappend_11170)) of - Some (stringappend_11180,stringappend_11190) => - (let stringappend_11200 = - (string_drop stringappend_11170 stringappend_11190) in + Some (stringappend_11510,stringappend_11520) => + (let stringappend_11530 = + (string_drop stringappend_11500 stringappend_11520) in + if ((case ((sep_matches_prefix stringappend_11530)) of + Some (stringappend_11540,stringappend_11550) => + (let stringappend_11560 = + (string_drop stringappend_11530 stringappend_11550) in if ((case ((hex_bits_12_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11200 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11560 :: (( 12 Word.word * ii))option)) of - Some (stringappend_11210,stringappend_11220) => - if(((string_drop stringappend_11200 stringappend_11220)) = ('''')) then + Some (stringappend_11570,stringappend_11580) => + if(((string_drop stringappend_11560 stringappend_11580)) = ('''')) then True else False | None => False )) then @@ -14248,83 +14323,83 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_11050 = (string_drop stringappend_10760 ((string_length (''jalr'')))) in + (let stringappend_11410 = (string_drop stringappend_11120 ((string_length (''jalr'')))) in (case - (case ((spc_matches_prefix stringappend_11050)) of - Some (stringappend_11060,stringappend_11070) => - (stringappend_11060, stringappend_11070) + (case ((spc_matches_prefix stringappend_11410)) of + Some (stringappend_11420,stringappend_11430) => + (stringappend_11420, stringappend_11430) ) of - (_, stringappend_11070) => - (let stringappend_11080 = (string_drop stringappend_11050 - stringappend_11070) in - (let (rd, stringappend_11100) = - ((case ((reg_name_matches_prefix stringappend_11080 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11090,stringappend_11100) => - (stringappend_11090, stringappend_11100) + (_, stringappend_11430) => + (let stringappend_11440 = (string_drop stringappend_11410 + stringappend_11430) in + (let (rd, stringappend_11460) = + ((case ((reg_name_matches_prefix stringappend_11440 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_11450,stringappend_11460) => + (stringappend_11450, stringappend_11460) )) in - (let stringappend_11110 = (string_drop stringappend_11080 - stringappend_11100) in + (let stringappend_11470 = (string_drop stringappend_11440 + stringappend_11460) in (case - (case ((sep_matches_prefix stringappend_11110)) of - Some (stringappend_11120,stringappend_11130) => - (stringappend_11120, stringappend_11130) + (case ((sep_matches_prefix stringappend_11470)) of + Some (stringappend_11480,stringappend_11490) => + (stringappend_11480, stringappend_11490) ) of - (_, stringappend_11130) => - (let stringappend_11140 = (string_drop stringappend_11110 - stringappend_11130) in - (let (rs1, stringappend_11160) = - ((case ((reg_name_matches_prefix stringappend_11140 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11150,stringappend_11160) => - (stringappend_11150, stringappend_11160) + (_, stringappend_11490) => + (let stringappend_11500 = (string_drop stringappend_11470 + stringappend_11490) in + (let (rs1, stringappend_11520) = + ((case ((reg_name_matches_prefix stringappend_11500 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_11510,stringappend_11520) => + (stringappend_11510, stringappend_11520) )) in - (let stringappend_11170 = (string_drop stringappend_11140 - stringappend_11160) in + (let stringappend_11530 = (string_drop stringappend_11500 + stringappend_11520) in (case - (case ((sep_matches_prefix stringappend_11170)) of - Some (stringappend_11180,stringappend_11190) => - (stringappend_11180, stringappend_11190) + (case ((sep_matches_prefix stringappend_11530)) of + Some (stringappend_11540,stringappend_11550) => + (stringappend_11540, stringappend_11550) ) of - (_, stringappend_11190) => - (let stringappend_11200 = (string_drop stringappend_11170 - stringappend_11190) in - (let (imm, stringappend_11220) = + (_, stringappend_11550) => + (let stringappend_11560 = (string_drop stringappend_11530 + stringappend_11550) in + (let (imm, stringappend_11580) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_11200 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_11210,stringappend_11220) => - (stringappend_11210, stringappend_11220) + stringappend_11560 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_11570,stringappend_11580) => + (stringappend_11570, stringappend_11580) )) in - if(((string_drop stringappend_11200 stringappend_11220)) = ('''')) then + if(((string_drop stringappend_11560 stringappend_11580)) = ('''')) then (RISCV_JALR (imm,rs1,rd)) else undefined)) )))) )))) )) - else if ((case ((btype_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_11240,stringappend_11250) => - (let stringappend_11260 = (string_drop stringappend_10760 stringappend_11250) in - if ((case ((spc_matches_prefix stringappend_11260)) of - Some (stringappend_11270,stringappend_11280) => - (let stringappend_11290 = (string_drop stringappend_11260 stringappend_11280) in - if ((case ((reg_name_matches_prefix stringappend_11290 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11300,stringappend_11310) => - (let stringappend_11320 = (string_drop stringappend_11290 stringappend_11310) in - if ((case ((sep_matches_prefix stringappend_11320)) of - Some (stringappend_11330,stringappend_11340) => - (let stringappend_11350 = (string_drop stringappend_11320 stringappend_11340) in - if ((case ((reg_name_matches_prefix stringappend_11350 + else if ((case ((btype_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_11600,stringappend_11610) => + (let stringappend_11620 = (string_drop stringappend_11120 stringappend_11610) in + if ((case ((spc_matches_prefix stringappend_11620)) of + Some (stringappend_11630,stringappend_11640) => + (let stringappend_11650 = (string_drop stringappend_11620 stringappend_11640) in + if ((case ((reg_name_matches_prefix stringappend_11650 :: (( 5 Word.word * ii))option)) of + Some (stringappend_11660,stringappend_11670) => + (let stringappend_11680 = (string_drop stringappend_11650 stringappend_11670) in + if ((case ((sep_matches_prefix stringappend_11680)) of + Some (stringappend_11690,stringappend_11700) => + (let stringappend_11710 = (string_drop stringappend_11680 stringappend_11700) in + if ((case ((reg_name_matches_prefix stringappend_11710 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11360,stringappend_11370) => - (let stringappend_11380 = - (string_drop stringappend_11350 stringappend_11370) in - if ((case ((sep_matches_prefix stringappend_11380)) of - Some (stringappend_11390,stringappend_11400) => - (let stringappend_11410 = - (string_drop stringappend_11380 stringappend_11400) in + Some (stringappend_11720,stringappend_11730) => + (let stringappend_11740 = + (string_drop stringappend_11710 stringappend_11730) in + if ((case ((sep_matches_prefix stringappend_11740)) of + Some (stringappend_11750,stringappend_11760) => + (let stringappend_11770 = + (string_drop stringappend_11740 stringappend_11760) in if ((case ((hex_bits_13_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11410 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11770 :: (( 13 Word.word * ii))option)) of - Some (stringappend_11420,stringappend_11430) => - if(((string_drop stringappend_11410 stringappend_11430)) = ('''')) then + Some (stringappend_11780,stringappend_11790) => + if(((string_drop stringappend_11770 stringappend_11790)) = ('''')) then True else False | None => False )) then @@ -14352,88 +14427,88 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where else False) | None => False )) then - (let (op1, stringappend_11250) = - ((case ((btype_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_11240,stringappend_11250) => - (stringappend_11240, stringappend_11250) + (let (op1, stringappend_11610) = + ((case ((btype_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_11600,stringappend_11610) => + (stringappend_11600, stringappend_11610) )) in - (let stringappend_11260 = (string_drop stringappend_10760 stringappend_11250) in + (let stringappend_11620 = (string_drop stringappend_11120 stringappend_11610) in (case - (case ((spc_matches_prefix stringappend_11260)) of - Some (stringappend_11270,stringappend_11280) => - (stringappend_11270, stringappend_11280) + (case ((spc_matches_prefix stringappend_11620)) of + Some (stringappend_11630,stringappend_11640) => + (stringappend_11630, stringappend_11640) ) of - (_, stringappend_11280) => - (let stringappend_11290 = (string_drop stringappend_11260 - stringappend_11280) in - (let (rs1, stringappend_11310) = - ((case ((reg_name_matches_prefix stringappend_11290 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11300,stringappend_11310) => - (stringappend_11300, stringappend_11310) + (_, stringappend_11640) => + (let stringappend_11650 = (string_drop stringappend_11620 + stringappend_11640) in + (let (rs1, stringappend_11670) = + ((case ((reg_name_matches_prefix stringappend_11650 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_11660,stringappend_11670) => + (stringappend_11660, stringappend_11670) )) in - (let stringappend_11320 = (string_drop stringappend_11290 - stringappend_11310) in + (let stringappend_11680 = (string_drop stringappend_11650 + stringappend_11670) in (case - (case ((sep_matches_prefix stringappend_11320)) of - Some (stringappend_11330,stringappend_11340) => - (stringappend_11330, stringappend_11340) + (case ((sep_matches_prefix stringappend_11680)) of + Some (stringappend_11690,stringappend_11700) => + (stringappend_11690, stringappend_11700) ) of - (_, stringappend_11340) => - (let stringappend_11350 = (string_drop stringappend_11320 - stringappend_11340) in - (let (rs2, stringappend_11370) = - ((case ((reg_name_matches_prefix stringappend_11350 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11360,stringappend_11370) => - (stringappend_11360, stringappend_11370) + (_, stringappend_11700) => + (let stringappend_11710 = (string_drop stringappend_11680 + stringappend_11700) in + (let (rs2, stringappend_11730) = + ((case ((reg_name_matches_prefix stringappend_11710 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_11720,stringappend_11730) => + (stringappend_11720, stringappend_11730) )) in - (let stringappend_11380 = (string_drop stringappend_11350 - stringappend_11370) in + (let stringappend_11740 = (string_drop stringappend_11710 + stringappend_11730) in (case - (case ((sep_matches_prefix stringappend_11380)) of - Some (stringappend_11390,stringappend_11400) => - (stringappend_11390, stringappend_11400) + (case ((sep_matches_prefix stringappend_11740)) of + Some (stringappend_11750,stringappend_11760) => + (stringappend_11750, stringappend_11760) ) of - (_, stringappend_11400) => - (let stringappend_11410 = (string_drop stringappend_11380 - stringappend_11400) in - (let (imm, stringappend_11430) = + (_, stringappend_11760) => + (let stringappend_11770 = (string_drop stringappend_11740 + stringappend_11760) in + (let (imm, stringappend_11790) = ((case ((hex_bits_13_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_11410 :: (( 13 Word.word * ii)) option)) of - Some (stringappend_11420,stringappend_11430) => - (stringappend_11420, stringappend_11430) + stringappend_11770 :: (( 13 Word.word * ii)) option)) of + Some (stringappend_11780,stringappend_11790) => + (stringappend_11780, stringappend_11790) )) in - if(((string_drop stringappend_11410 stringappend_11430)) = ('''')) then + if(((string_drop stringappend_11770 stringappend_11790)) = ('''')) then (BTYPE (imm,rs2,rs1,op1)) else undefined)) )))) )))) ))) - else if ((case ((itype_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_11450,stringappend_11460) => - (let stringappend_11470 = (string_drop stringappend_10760 stringappend_11460) in - if ((case ((spc_matches_prefix stringappend_11470)) of - Some (stringappend_11480,stringappend_11490) => - (let stringappend_11500 = (string_drop stringappend_11470 stringappend_11490) in - if ((case ((reg_name_matches_prefix stringappend_11500 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11510,stringappend_11520) => - (let stringappend_11530 = (string_drop stringappend_11500 stringappend_11520) in - if ((case ((sep_matches_prefix stringappend_11530)) of - Some (stringappend_11540,stringappend_11550) => - (let stringappend_11560 = (string_drop stringappend_11530 stringappend_11550) in - if ((case ((reg_name_matches_prefix stringappend_11560 + else if ((case ((itype_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_11810,stringappend_11820) => + (let stringappend_11830 = (string_drop stringappend_11120 stringappend_11820) in + if ((case ((spc_matches_prefix stringappend_11830)) of + Some (stringappend_11840,stringappend_11850) => + (let stringappend_11860 = (string_drop stringappend_11830 stringappend_11850) in + if ((case ((reg_name_matches_prefix stringappend_11860 :: (( 5 Word.word * ii))option)) of + Some (stringappend_11870,stringappend_11880) => + (let stringappend_11890 = (string_drop stringappend_11860 stringappend_11880) in + if ((case ((sep_matches_prefix stringappend_11890)) of + Some (stringappend_11900,stringappend_11910) => + (let stringappend_11920 = (string_drop stringappend_11890 stringappend_11910) in + if ((case ((reg_name_matches_prefix stringappend_11920 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11570,stringappend_11580) => - (let stringappend_11590 = - (string_drop stringappend_11560 stringappend_11580) in - if ((case ((sep_matches_prefix stringappend_11590)) of - Some (stringappend_11600,stringappend_11610) => - (let stringappend_11620 = - (string_drop stringappend_11590 stringappend_11610) in + Some (stringappend_11930,stringappend_11940) => + (let stringappend_11950 = + (string_drop stringappend_11920 stringappend_11940) in + if ((case ((sep_matches_prefix stringappend_11950)) of + Some (stringappend_11960,stringappend_11970) => + (let stringappend_11980 = + (string_drop stringappend_11950 stringappend_11970) in if ((case ((hex_bits_12_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11620 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11980 :: (( 12 Word.word * ii))option)) of - Some (stringappend_11630,stringappend_11640) => - if(((string_drop stringappend_11620 stringappend_11640)) = ('''')) then + Some (stringappend_11990,stringappend_12000) => + if(((string_drop stringappend_11980 stringappend_12000)) = ('''')) then True else False | None => False )) then @@ -14461,84 +14536,84 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where else False) | None => False )) then - (let (op1, stringappend_11460) = - ((case ((itype_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_11450,stringappend_11460) => - (stringappend_11450, stringappend_11460) + (let (op1, stringappend_11820) = + ((case ((itype_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_11810,stringappend_11820) => + (stringappend_11810, stringappend_11820) )) in - (let stringappend_11470 = (string_drop stringappend_10760 stringappend_11460) in + (let stringappend_11830 = (string_drop stringappend_11120 stringappend_11820) in (case - (case ((spc_matches_prefix stringappend_11470)) of - Some (stringappend_11480,stringappend_11490) => - (stringappend_11480, stringappend_11490) + (case ((spc_matches_prefix stringappend_11830)) of + Some (stringappend_11840,stringappend_11850) => + (stringappend_11840, stringappend_11850) ) of - (_, stringappend_11490) => - (let stringappend_11500 = (string_drop stringappend_11470 - stringappend_11490) in - (let (rd, stringappend_11520) = - ((case ((reg_name_matches_prefix stringappend_11500 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11510,stringappend_11520) => - (stringappend_11510, stringappend_11520) + (_, stringappend_11850) => + (let stringappend_11860 = (string_drop stringappend_11830 + stringappend_11850) in + (let (rd, stringappend_11880) = + ((case ((reg_name_matches_prefix stringappend_11860 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_11870,stringappend_11880) => + (stringappend_11870, stringappend_11880) )) in - (let stringappend_11530 = (string_drop stringappend_11500 - stringappend_11520) in + (let stringappend_11890 = (string_drop stringappend_11860 + stringappend_11880) in (case - (case ((sep_matches_prefix stringappend_11530)) of - Some (stringappend_11540,stringappend_11550) => - (stringappend_11540, stringappend_11550) + (case ((sep_matches_prefix stringappend_11890)) of + Some (stringappend_11900,stringappend_11910) => + (stringappend_11900, stringappend_11910) ) of - (_, stringappend_11550) => - (let stringappend_11560 = (string_drop stringappend_11530 - stringappend_11550) in - (let (rs1, stringappend_11580) = - ((case ((reg_name_matches_prefix stringappend_11560 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11570,stringappend_11580) => - (stringappend_11570, stringappend_11580) + (_, stringappend_11910) => + (let stringappend_11920 = (string_drop stringappend_11890 + stringappend_11910) in + (let (rs1, stringappend_11940) = + ((case ((reg_name_matches_prefix stringappend_11920 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_11930,stringappend_11940) => + (stringappend_11930, stringappend_11940) )) in - (let stringappend_11590 = (string_drop stringappend_11560 - stringappend_11580) in + (let stringappend_11950 = (string_drop stringappend_11920 + stringappend_11940) in (case - (case ((sep_matches_prefix stringappend_11590)) of - Some (stringappend_11600,stringappend_11610) => - (stringappend_11600, stringappend_11610) + (case ((sep_matches_prefix stringappend_11950)) of + Some (stringappend_11960,stringappend_11970) => + (stringappend_11960, stringappend_11970) ) of - (_, stringappend_11610) => - (let stringappend_11620 = (string_drop stringappend_11590 - stringappend_11610) in - (let (imm, stringappend_11640) = + (_, stringappend_11970) => + (let stringappend_11980 = (string_drop stringappend_11950 + stringappend_11970) in + (let (imm, stringappend_12000) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_11620 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_11630,stringappend_11640) => - (stringappend_11630, stringappend_11640) + stringappend_11980 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_11990,stringappend_12000) => + (stringappend_11990, stringappend_12000) )) in - if(((string_drop stringappend_11620 stringappend_11640)) = ('''')) then + if(((string_drop stringappend_11980 stringappend_12000)) = ('''')) then (ITYPE (imm,rs1,rd,op1)) else undefined)) )))) )))) ))) - else if ((case ((shiftiop_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_11660,stringappend_11670) => - (let stringappend_11680 = (string_drop stringappend_10760 stringappend_11670) in - if ((case ((spc_matches_prefix stringappend_11680)) of - Some (stringappend_11690,stringappend_11700) => - (let stringappend_11710 = (string_drop stringappend_11680 stringappend_11700) in - if ((case ((reg_name_matches_prefix stringappend_11710 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11720,stringappend_11730) => - (let stringappend_11740 = (string_drop stringappend_11710 stringappend_11730) in - if ((case ((sep_matches_prefix stringappend_11740)) of - Some (stringappend_11750,stringappend_11760) => - (let stringappend_11770 = (string_drop stringappend_11740 stringappend_11760) in - if ((case ((reg_name_matches_prefix stringappend_11770 + else if ((case ((shiftiop_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_12020,stringappend_12030) => + (let stringappend_12040 = (string_drop stringappend_11120 stringappend_12030) in + if ((case ((spc_matches_prefix stringappend_12040)) of + Some (stringappend_12050,stringappend_12060) => + (let stringappend_12070 = (string_drop stringappend_12040 stringappend_12060) in + if ((case ((reg_name_matches_prefix stringappend_12070 :: (( 5 Word.word * ii))option)) of + Some (stringappend_12080,stringappend_12090) => + (let stringappend_12100 = (string_drop stringappend_12070 stringappend_12090) in + if ((case ((sep_matches_prefix stringappend_12100)) of + Some (stringappend_12110,stringappend_12120) => + (let stringappend_12130 = (string_drop stringappend_12100 stringappend_12120) in + if ((case ((reg_name_matches_prefix stringappend_12130 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11780,stringappend_11790) => - (let stringappend_11800 = - (string_drop stringappend_11770 stringappend_11790) in + Some (stringappend_12140,stringappend_12150) => + (let stringappend_12160 = + (string_drop stringappend_12130 stringappend_12150) in if ((case ((hex_bits_6_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_11800 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_12160 :: (( 6 Word.word * ii))option)) of - Some (stringappend_11810,stringappend_11820) => - if(((string_drop stringappend_11800 stringappend_11820)) = ('''')) then + Some (stringappend_12170,stringappend_12180) => + if(((string_drop stringappend_12160 stringappend_12180)) = ('''')) then True else False | None => False )) then @@ -14562,78 +14637,78 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where else False) | None => False )) then - (let (op1, stringappend_11670) = - ((case ((shiftiop_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_11660,stringappend_11670) => - (stringappend_11660, stringappend_11670) + (let (op1, stringappend_12030) = + ((case ((shiftiop_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_12020,stringappend_12030) => + (stringappend_12020, stringappend_12030) )) in - (let stringappend_11680 = (string_drop stringappend_10760 stringappend_11670) in + (let stringappend_12040 = (string_drop stringappend_11120 stringappend_12030) in (case - (case ((spc_matches_prefix stringappend_11680)) of - Some (stringappend_11690,stringappend_11700) => - (stringappend_11690, stringappend_11700) + (case ((spc_matches_prefix stringappend_12040)) of + Some (stringappend_12050,stringappend_12060) => + (stringappend_12050, stringappend_12060) ) of - (_, stringappend_11700) => - (let stringappend_11710 = (string_drop stringappend_11680 - stringappend_11700) in - (let (rd, stringappend_11730) = - ((case ((reg_name_matches_prefix stringappend_11710 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11720,stringappend_11730) => - (stringappend_11720, stringappend_11730) + (_, stringappend_12060) => + (let stringappend_12070 = (string_drop stringappend_12040 + stringappend_12060) in + (let (rd, stringappend_12090) = + ((case ((reg_name_matches_prefix stringappend_12070 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12080,stringappend_12090) => + (stringappend_12080, stringappend_12090) )) in - (let stringappend_11740 = (string_drop stringappend_11710 - stringappend_11730) in + (let stringappend_12100 = (string_drop stringappend_12070 + stringappend_12090) in (case - (case ((sep_matches_prefix stringappend_11740)) of - Some (stringappend_11750,stringappend_11760) => - (stringappend_11750, stringappend_11760) + (case ((sep_matches_prefix stringappend_12100)) of + Some (stringappend_12110,stringappend_12120) => + (stringappend_12110, stringappend_12120) ) of - (_, stringappend_11760) => - (let stringappend_11770 = (string_drop stringappend_11740 - stringappend_11760) in - (let (rs1, stringappend_11790) = - ((case ((reg_name_matches_prefix stringappend_11770 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11780,stringappend_11790) => - (stringappend_11780, stringappend_11790) + (_, stringappend_12120) => + (let stringappend_12130 = (string_drop stringappend_12100 + stringappend_12120) in + (let (rs1, stringappend_12150) = + ((case ((reg_name_matches_prefix stringappend_12130 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12140,stringappend_12150) => + (stringappend_12140, stringappend_12150) )) in - (let stringappend_11800 = (string_drop stringappend_11770 - stringappend_11790) in - (let (shamt, stringappend_11820) = + (let stringappend_12160 = (string_drop stringappend_12130 + stringappend_12150) in + (let (shamt, stringappend_12180) = ((case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_11800 :: (( 6 Word.word * ii)) option)) of - Some (stringappend_11810,stringappend_11820) => - (stringappend_11810, stringappend_11820) + stringappend_12160 :: (( 6 Word.word * ii)) option)) of + Some (stringappend_12170,stringappend_12180) => + (stringappend_12170, stringappend_12180) )) in - if(((string_drop stringappend_11800 stringappend_11820)) = ('''')) then + if(((string_drop stringappend_12160 stringappend_12180)) = ('''')) then (SHIFTIOP (shamt,rs1,rd,op1)) else undefined)))) )))) ))) - else if ((case ((rtype_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_11840,stringappend_11850) => - (let stringappend_11860 = (string_drop stringappend_10760 stringappend_11850) in - if ((case ((spc_matches_prefix stringappend_11860)) of - Some (stringappend_11870,stringappend_11880) => - (let stringappend_11890 = (string_drop stringappend_11860 stringappend_11880) in - if ((case ((reg_name_matches_prefix stringappend_11890 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11900,stringappend_11910) => - (let stringappend_11920 = (string_drop stringappend_11890 stringappend_11910) in - if ((case ((sep_matches_prefix stringappend_11920)) of - Some (stringappend_11930,stringappend_11940) => - (let stringappend_11950 = (string_drop stringappend_11920 stringappend_11940) in - if ((case ((reg_name_matches_prefix stringappend_11950 + else if ((case ((rtype_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_12200,stringappend_12210) => + (let stringappend_12220 = (string_drop stringappend_11120 stringappend_12210) in + if ((case ((spc_matches_prefix stringappend_12220)) of + Some (stringappend_12230,stringappend_12240) => + (let stringappend_12250 = (string_drop stringappend_12220 stringappend_12240) in + if ((case ((reg_name_matches_prefix stringappend_12250 :: (( 5 Word.word * ii))option)) of + Some (stringappend_12260,stringappend_12270) => + (let stringappend_12280 = (string_drop stringappend_12250 stringappend_12270) in + if ((case ((sep_matches_prefix stringappend_12280)) of + Some (stringappend_12290,stringappend_12300) => + (let stringappend_12310 = (string_drop stringappend_12280 stringappend_12300) in + if ((case ((reg_name_matches_prefix stringappend_12310 :: (( 5 Word.word * ii))option)) of - Some (stringappend_11960,stringappend_11970) => - (let stringappend_11980 = - (string_drop stringappend_11950 stringappend_11970) in - if ((case ((sep_matches_prefix stringappend_11980)) of - Some (stringappend_11990,stringappend_12000) => - (let stringappend_12010 = - (string_drop stringappend_11980 stringappend_12000) in - if ((case ((reg_name_matches_prefix stringappend_12010 + Some (stringappend_12320,stringappend_12330) => + (let stringappend_12340 = + (string_drop stringappend_12310 stringappend_12330) in + if ((case ((sep_matches_prefix stringappend_12340)) of + Some (stringappend_12350,stringappend_12360) => + (let stringappend_12370 = + (string_drop stringappend_12340 stringappend_12360) in + if ((case ((reg_name_matches_prefix stringappend_12370 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12020,stringappend_12030) => - if(((string_drop stringappend_12010 stringappend_12030)) = ('''')) then + Some (stringappend_12380,stringappend_12390) => + if(((string_drop stringappend_12370 stringappend_12390)) = ('''')) then True else False | None => False )) then @@ -14661,111 +14736,111 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where else False) | None => False )) then - (let (op1, stringappend_11850) = - ((case ((rtype_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_11840,stringappend_11850) => - (stringappend_11840, stringappend_11850) + (let (op1, stringappend_12210) = + ((case ((rtype_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_12200,stringappend_12210) => + (stringappend_12200, stringappend_12210) )) in - (let stringappend_11860 = (string_drop stringappend_10760 stringappend_11850) in + (let stringappend_12220 = (string_drop stringappend_11120 stringappend_12210) in (case - (case ((spc_matches_prefix stringappend_11860)) of - Some (stringappend_11870,stringappend_11880) => - (stringappend_11870, stringappend_11880) + (case ((spc_matches_prefix stringappend_12220)) of + Some (stringappend_12230,stringappend_12240) => + (stringappend_12230, stringappend_12240) ) of - (_, stringappend_11880) => - (let stringappend_11890 = (string_drop stringappend_11860 - stringappend_11880) in - (let (rd, stringappend_11910) = - ((case ((reg_name_matches_prefix stringappend_11890 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11900,stringappend_11910) => - (stringappend_11900, stringappend_11910) + (_, stringappend_12240) => + (let stringappend_12250 = (string_drop stringappend_12220 + stringappend_12240) in + (let (rd, stringappend_12270) = + ((case ((reg_name_matches_prefix stringappend_12250 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12260,stringappend_12270) => + (stringappend_12260, stringappend_12270) )) in - (let stringappend_11920 = (string_drop stringappend_11890 - stringappend_11910) in + (let stringappend_12280 = (string_drop stringappend_12250 + stringappend_12270) in (case - (case ((sep_matches_prefix stringappend_11920)) of - Some (stringappend_11930,stringappend_11940) => - (stringappend_11930, stringappend_11940) + (case ((sep_matches_prefix stringappend_12280)) of + Some (stringappend_12290,stringappend_12300) => + (stringappend_12290, stringappend_12300) ) of - (_, stringappend_11940) => - (let stringappend_11950 = (string_drop stringappend_11920 - stringappend_11940) in - (let (rs1, stringappend_11970) = - ((case ((reg_name_matches_prefix stringappend_11950 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_11960,stringappend_11970) => - (stringappend_11960, stringappend_11970) + (_, stringappend_12300) => + (let stringappend_12310 = (string_drop stringappend_12280 + stringappend_12300) in + (let (rs1, stringappend_12330) = + ((case ((reg_name_matches_prefix stringappend_12310 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12320,stringappend_12330) => + (stringappend_12320, stringappend_12330) )) in - (let stringappend_11980 = (string_drop stringappend_11950 - stringappend_11970) in + (let stringappend_12340 = (string_drop stringappend_12310 + stringappend_12330) in (case - (case ((sep_matches_prefix stringappend_11980)) of - Some (stringappend_11990,stringappend_12000) => - (stringappend_11990, stringappend_12000) + (case ((sep_matches_prefix stringappend_12340)) of + Some (stringappend_12350,stringappend_12360) => + (stringappend_12350, stringappend_12360) ) of - (_, stringappend_12000) => - (let stringappend_12010 = (string_drop stringappend_11980 - stringappend_12000) in - (let (rs2, stringappend_12030) = - ((case ((reg_name_matches_prefix stringappend_12010 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12020,stringappend_12030) => - (stringappend_12020, stringappend_12030) + (_, stringappend_12360) => + (let stringappend_12370 = (string_drop stringappend_12340 + stringappend_12360) in + (let (rs2, stringappend_12390) = + ((case ((reg_name_matches_prefix stringappend_12370 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12380,stringappend_12390) => + (stringappend_12380, stringappend_12390) )) in - if(((string_drop stringappend_12010 stringappend_12030)) = ('''')) then + if(((string_drop stringappend_12370 stringappend_12390)) = ('''')) then (RTYPE (rs2,rs1,rd,op1)) else undefined)) )))) )))) ))) - else if (((((string_startswith stringappend_10760 (''l''))) \<and> ((let stringappend_12050 = (string_drop stringappend_10760 ((string_length (''l'')))) in - if ((case ((size_mnemonic_matches_prefix stringappend_12050)) of - Some (stringappend_12060,stringappend_12070) => - (let stringappend_12080 = (string_drop stringappend_12050 stringappend_12070) in - if ((case ((maybe_u_matches_prefix stringappend_12080)) of - Some (stringappend_12090,stringappend_12100) => - (let stringappend_12110 = - (string_drop stringappend_12080 stringappend_12100) in - if ((case ((maybe_aq_matches_prefix stringappend_12110)) of - Some (stringappend_12120,stringappend_12130) => - (let stringappend_12140 = - (string_drop stringappend_12110 stringappend_12130) in - if ((case ((maybe_rl_matches_prefix stringappend_12140)) of - Some (stringappend_12150,stringappend_12160) => - (let stringappend_12170 = - (string_drop stringappend_12140 stringappend_12160) in - if ((case ((spc_matches_prefix stringappend_12170)) of - Some (stringappend_12180,stringappend_12190) => - (let stringappend_12200 = - (string_drop stringappend_12170 stringappend_12190) in - if ((case ((reg_name_matches_prefix stringappend_12200 + else if (((((string_startswith stringappend_11120 (''l''))) \<and> ((let stringappend_12410 = (string_drop stringappend_11120 ((string_length (''l'')))) in + if ((case ((size_mnemonic_matches_prefix stringappend_12410)) of + Some (stringappend_12420,stringappend_12430) => + (let stringappend_12440 = (string_drop stringappend_12410 stringappend_12430) in + if ((case ((maybe_u_matches_prefix stringappend_12440)) of + Some (stringappend_12450,stringappend_12460) => + (let stringappend_12470 = + (string_drop stringappend_12440 stringappend_12460) in + if ((case ((maybe_aq_matches_prefix stringappend_12470)) of + Some (stringappend_12480,stringappend_12490) => + (let stringappend_12500 = + (string_drop stringappend_12470 stringappend_12490) in + if ((case ((maybe_rl_matches_prefix stringappend_12500)) of + Some (stringappend_12510,stringappend_12520) => + (let stringappend_12530 = + (string_drop stringappend_12500 stringappend_12520) in + if ((case ((spc_matches_prefix stringappend_12530)) of + Some (stringappend_12540,stringappend_12550) => + (let stringappend_12560 = + (string_drop stringappend_12530 stringappend_12550) in + if ((case ((reg_name_matches_prefix stringappend_12560 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12210,stringappend_12220) => - (let stringappend_12230 = - (string_drop stringappend_12200 stringappend_12220) in - if ((case ((sep_matches_prefix stringappend_12230)) of - Some (stringappend_12240,stringappend_12250) => - (let stringappend_12260 = - (string_drop stringappend_12230 - stringappend_12250) in + Some (stringappend_12570,stringappend_12580) => + (let stringappend_12590 = + (string_drop stringappend_12560 stringappend_12580) in + if ((case ((sep_matches_prefix stringappend_12590)) of + Some (stringappend_12600,stringappend_12610) => + (let stringappend_12620 = + (string_drop stringappend_12590 + stringappend_12610) in if ((case ((reg_name_matches_prefix - stringappend_12260 + stringappend_12620 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12270,stringappend_12280) => - (let stringappend_12290 = - (string_drop stringappend_12260 - stringappend_12280) in + Some (stringappend_12630,stringappend_12640) => + (let stringappend_12650 = + (string_drop stringappend_12620 + stringappend_12640) in if ((case ((sep_matches_prefix - stringappend_12290)) of + stringappend_12650)) of Some - (stringappend_12300,stringappend_12310) => - (let stringappend_12320 = - (string_drop stringappend_12290 - stringappend_12310) in + (stringappend_12660,stringappend_12670) => + (let stringappend_12680 = + (string_drop stringappend_12650 + stringappend_12670) in if ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_12320 + stringappend_12680 :: (( 12 Word.word * ii))option)) of Some - (stringappend_12330,stringappend_12340) => - if(((string_drop stringappend_12320 stringappend_12340)) = ('''')) then + (stringappend_12690,stringappend_12700) => + if(((string_drop stringappend_12680 stringappend_12700)) = ('''')) then True else False | None => False )) then @@ -14807,124 +14882,124 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_12050 = (string_drop stringappend_10760 ((string_length (''l'')))) in - (let (size1, stringappend_12070) = - ((case ((size_mnemonic_matches_prefix stringappend_12050)) of - Some (stringappend_12060,stringappend_12070) => - (stringappend_12060, stringappend_12070) + (let stringappend_12410 = (string_drop stringappend_11120 ((string_length (''l'')))) in + (let (size1, stringappend_12430) = + ((case ((size_mnemonic_matches_prefix stringappend_12410)) of + Some (stringappend_12420,stringappend_12430) => + (stringappend_12420, stringappend_12430) )) in - (let stringappend_12080 = (string_drop stringappend_12050 stringappend_12070) in - (let (is_unsigned, stringappend_12100) = - ((case ((maybe_u_matches_prefix stringappend_12080)) of - Some (stringappend_12090,stringappend_12100) => - (stringappend_12090, stringappend_12100) + (let stringappend_12440 = (string_drop stringappend_12410 stringappend_12430) in + (let (is_unsigned, stringappend_12460) = + ((case ((maybe_u_matches_prefix stringappend_12440)) of + Some (stringappend_12450,stringappend_12460) => + (stringappend_12450, stringappend_12460) )) in - (let stringappend_12110 = (string_drop stringappend_12080 stringappend_12100) in - (let (aq, stringappend_12130) = - ((case ((maybe_aq_matches_prefix stringappend_12110)) of - Some (stringappend_12120,stringappend_12130) => - (stringappend_12120, stringappend_12130) + (let stringappend_12470 = (string_drop stringappend_12440 stringappend_12460) in + (let (aq, stringappend_12490) = + ((case ((maybe_aq_matches_prefix stringappend_12470)) of + Some (stringappend_12480,stringappend_12490) => + (stringappend_12480, stringappend_12490) )) in - (let stringappend_12140 = (string_drop stringappend_12110 stringappend_12130) in - (let (rl, stringappend_12160) = - ((case ((maybe_rl_matches_prefix stringappend_12140)) of - Some (stringappend_12150,stringappend_12160) => - (stringappend_12150, stringappend_12160) + (let stringappend_12500 = (string_drop stringappend_12470 stringappend_12490) in + (let (rl, stringappend_12520) = + ((case ((maybe_rl_matches_prefix stringappend_12500)) of + Some (stringappend_12510,stringappend_12520) => + (stringappend_12510, stringappend_12520) )) in - (let stringappend_12170 = (string_drop stringappend_12140 stringappend_12160) in + (let stringappend_12530 = (string_drop stringappend_12500 stringappend_12520) in (case - (case ((spc_matches_prefix stringappend_12170)) of - Some (stringappend_12180,stringappend_12190) => - (stringappend_12180, stringappend_12190) + (case ((spc_matches_prefix stringappend_12530)) of + Some (stringappend_12540,stringappend_12550) => + (stringappend_12540, stringappend_12550) ) of - (_, stringappend_12190) => - (let stringappend_12200 = (string_drop stringappend_12170 - stringappend_12190) in - (let (rd, stringappend_12220) = - ((case ((reg_name_matches_prefix stringappend_12200 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12210,stringappend_12220) => - (stringappend_12210, stringappend_12220) + (_, stringappend_12550) => + (let stringappend_12560 = (string_drop stringappend_12530 + stringappend_12550) in + (let (rd, stringappend_12580) = + ((case ((reg_name_matches_prefix stringappend_12560 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12570,stringappend_12580) => + (stringappend_12570, stringappend_12580) )) in - (let stringappend_12230 = (string_drop stringappend_12200 - stringappend_12220) in + (let stringappend_12590 = (string_drop stringappend_12560 + stringappend_12580) in (case - (case ((sep_matches_prefix stringappend_12230)) of - Some (stringappend_12240,stringappend_12250) => - (stringappend_12240, stringappend_12250) + (case ((sep_matches_prefix stringappend_12590)) of + Some (stringappend_12600,stringappend_12610) => + (stringappend_12600, stringappend_12610) ) of - (_, stringappend_12250) => - (let stringappend_12260 = (string_drop stringappend_12230 - stringappend_12250) in - (let (rs1, stringappend_12280) = - ((case ((reg_name_matches_prefix stringappend_12260 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12270,stringappend_12280) => - (stringappend_12270, stringappend_12280) + (_, stringappend_12610) => + (let stringappend_12620 = (string_drop stringappend_12590 + stringappend_12610) in + (let (rs1, stringappend_12640) = + ((case ((reg_name_matches_prefix stringappend_12620 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12630,stringappend_12640) => + (stringappend_12630, stringappend_12640) )) in - (let stringappend_12290 = (string_drop stringappend_12260 - stringappend_12280) in + (let stringappend_12650 = (string_drop stringappend_12620 + stringappend_12640) in (case - (case ((sep_matches_prefix stringappend_12290)) of - Some (stringappend_12300,stringappend_12310) => - (stringappend_12300, stringappend_12310) + (case ((sep_matches_prefix stringappend_12650)) of + Some (stringappend_12660,stringappend_12670) => + (stringappend_12660, stringappend_12670) ) of - (_, stringappend_12310) => - (let stringappend_12320 = (string_drop stringappend_12290 - stringappend_12310) in - (let (imm, stringappend_12340) = + (_, stringappend_12670) => + (let stringappend_12680 = (string_drop stringappend_12650 + stringappend_12670) in + (let (imm, stringappend_12700) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_12320 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_12330,stringappend_12340) => - (stringappend_12330, stringappend_12340) + stringappend_12680 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_12690,stringappend_12700) => + (stringappend_12690, stringappend_12700) )) in - if(((string_drop stringappend_12320 stringappend_12340)) = ('''')) then + if(((string_drop stringappend_12680 stringappend_12700)) = ('''')) then (LOAD (imm,rs1,rd,is_unsigned,size1,aq,rl)) else undefined)) )))) )))) )))))))))) - else if (((((string_startswith stringappend_10760 (''s''))) \<and> ((let stringappend_12360 = (string_drop stringappend_10760 ((string_length (''s'')))) in - if ((case ((size_mnemonic_matches_prefix stringappend_12360)) of - Some (stringappend_12370,stringappend_12380) => - (let stringappend_12390 = (string_drop stringappend_12360 stringappend_12380) in - if ((case ((maybe_aq_matches_prefix stringappend_12390)) of - Some (stringappend_12400,stringappend_12410) => - (let stringappend_12420 = - (string_drop stringappend_12390 stringappend_12410) in - if ((case ((maybe_rl_matches_prefix stringappend_12420)) of - Some (stringappend_12430,stringappend_12440) => - (let stringappend_12450 = - (string_drop stringappend_12420 stringappend_12440) in - if ((case ((spc_matches_prefix stringappend_12450)) of - Some (stringappend_12460,stringappend_12470) => - (let stringappend_12480 = - (string_drop stringappend_12450 stringappend_12470) in - if ((case ((reg_name_matches_prefix stringappend_12480 + else if (((((string_startswith stringappend_11120 (''s''))) \<and> ((let stringappend_12720 = (string_drop stringappend_11120 ((string_length (''s'')))) in + if ((case ((size_mnemonic_matches_prefix stringappend_12720)) of + Some (stringappend_12730,stringappend_12740) => + (let stringappend_12750 = (string_drop stringappend_12720 stringappend_12740) in + if ((case ((maybe_aq_matches_prefix stringappend_12750)) of + Some (stringappend_12760,stringappend_12770) => + (let stringappend_12780 = + (string_drop stringappend_12750 stringappend_12770) in + if ((case ((maybe_rl_matches_prefix stringappend_12780)) of + Some (stringappend_12790,stringappend_12800) => + (let stringappend_12810 = + (string_drop stringappend_12780 stringappend_12800) in + if ((case ((spc_matches_prefix stringappend_12810)) of + Some (stringappend_12820,stringappend_12830) => + (let stringappend_12840 = + (string_drop stringappend_12810 stringappend_12830) in + if ((case ((reg_name_matches_prefix stringappend_12840 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12490,stringappend_12500) => - (let stringappend_12510 = - (string_drop stringappend_12480 stringappend_12500) in - if ((case ((sep_matches_prefix stringappend_12510)) of - Some (stringappend_12520,stringappend_12530) => - (let stringappend_12540 = - (string_drop stringappend_12510 stringappend_12530) in - if ((case ((reg_name_matches_prefix stringappend_12540 + Some (stringappend_12850,stringappend_12860) => + (let stringappend_12870 = + (string_drop stringappend_12840 stringappend_12860) in + if ((case ((sep_matches_prefix stringappend_12870)) of + Some (stringappend_12880,stringappend_12890) => + (let stringappend_12900 = + (string_drop stringappend_12870 stringappend_12890) in + if ((case ((reg_name_matches_prefix stringappend_12900 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12550,stringappend_12560) => - (let stringappend_12570 = - (string_drop stringappend_12540 - stringappend_12560) in - if ((case ((sep_matches_prefix stringappend_12570)) of - Some (stringappend_12580,stringappend_12590) => - (let stringappend_12600 = - (string_drop stringappend_12570 - stringappend_12590) in + Some (stringappend_12910,stringappend_12920) => + (let stringappend_12930 = + (string_drop stringappend_12900 + stringappend_12920) in + if ((case ((sep_matches_prefix stringappend_12930)) of + Some (stringappend_12940,stringappend_12950) => + (let stringappend_12960 = + (string_drop stringappend_12930 + stringappend_12950) in if ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_12600 + stringappend_12960 :: (( 12 Word.word * ii))option)) of Some - (stringappend_12610,stringappend_12620) => - if(((string_drop stringappend_12600 stringappend_12620)) = ('''')) then + (stringappend_12970,stringappend_12980) => + if(((string_drop stringappend_12960 stringappend_12980)) = ('''')) then True else False | None => False )) then @@ -14962,102 +15037,102 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_12360 = (string_drop stringappend_10760 ((string_length (''s'')))) in - (let (size1, stringappend_12380) = - ((case ((size_mnemonic_matches_prefix stringappend_12360)) of - Some (stringappend_12370,stringappend_12380) => - (stringappend_12370, stringappend_12380) + (let stringappend_12720 = (string_drop stringappend_11120 ((string_length (''s'')))) in + (let (size1, stringappend_12740) = + ((case ((size_mnemonic_matches_prefix stringappend_12720)) of + Some (stringappend_12730,stringappend_12740) => + (stringappend_12730, stringappend_12740) )) in - (let stringappend_12390 = (string_drop stringappend_12360 stringappend_12380) in - (let (aq, stringappend_12410) = - ((case ((maybe_aq_matches_prefix stringappend_12390)) of - Some (stringappend_12400,stringappend_12410) => - (stringappend_12400, stringappend_12410) + (let stringappend_12750 = (string_drop stringappend_12720 stringappend_12740) in + (let (aq, stringappend_12770) = + ((case ((maybe_aq_matches_prefix stringappend_12750)) of + Some (stringappend_12760,stringappend_12770) => + (stringappend_12760, stringappend_12770) )) in - (let stringappend_12420 = (string_drop stringappend_12390 stringappend_12410) in - (let (rl, stringappend_12440) = - ((case ((maybe_rl_matches_prefix stringappend_12420)) of - Some (stringappend_12430,stringappend_12440) => - (stringappend_12430, stringappend_12440) + (let stringappend_12780 = (string_drop stringappend_12750 stringappend_12770) in + (let (rl, stringappend_12800) = + ((case ((maybe_rl_matches_prefix stringappend_12780)) of + Some (stringappend_12790,stringappend_12800) => + (stringappend_12790, stringappend_12800) )) in - (let stringappend_12450 = (string_drop stringappend_12420 stringappend_12440) in + (let stringappend_12810 = (string_drop stringappend_12780 stringappend_12800) in (case - (case ((spc_matches_prefix stringappend_12450)) of - Some (stringappend_12460,stringappend_12470) => - (stringappend_12460, stringappend_12470) + (case ((spc_matches_prefix stringappend_12810)) of + Some (stringappend_12820,stringappend_12830) => + (stringappend_12820, stringappend_12830) ) of - (_, stringappend_12470) => - (let stringappend_12480 = (string_drop stringappend_12450 - stringappend_12470) in - (let (rd, stringappend_12500) = - ((case ((reg_name_matches_prefix stringappend_12480 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12490,stringappend_12500) => - (stringappend_12490, stringappend_12500) + (_, stringappend_12830) => + (let stringappend_12840 = (string_drop stringappend_12810 + stringappend_12830) in + (let (rd, stringappend_12860) = + ((case ((reg_name_matches_prefix stringappend_12840 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12850,stringappend_12860) => + (stringappend_12850, stringappend_12860) )) in - (let stringappend_12510 = (string_drop stringappend_12480 - stringappend_12500) in + (let stringappend_12870 = (string_drop stringappend_12840 + stringappend_12860) in (case - (case ((sep_matches_prefix stringappend_12510)) of - Some (stringappend_12520,stringappend_12530) => - (stringappend_12520, stringappend_12530) + (case ((sep_matches_prefix stringappend_12870)) of + Some (stringappend_12880,stringappend_12890) => + (stringappend_12880, stringappend_12890) ) of - (_, stringappend_12530) => - (let stringappend_12540 = (string_drop stringappend_12510 - stringappend_12530) in - (let (rs1, stringappend_12560) = - ((case ((reg_name_matches_prefix stringappend_12540 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12550,stringappend_12560) => - (stringappend_12550, stringappend_12560) + (_, stringappend_12890) => + (let stringappend_12900 = (string_drop stringappend_12870 + stringappend_12890) in + (let (rs1, stringappend_12920) = + ((case ((reg_name_matches_prefix stringappend_12900 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_12910,stringappend_12920) => + (stringappend_12910, stringappend_12920) )) in - (let stringappend_12570 = (string_drop stringappend_12540 - stringappend_12560) in + (let stringappend_12930 = (string_drop stringappend_12900 + stringappend_12920) in (case - (case ((sep_matches_prefix stringappend_12570)) of - Some (stringappend_12580,stringappend_12590) => - (stringappend_12580, stringappend_12590) + (case ((sep_matches_prefix stringappend_12930)) of + Some (stringappend_12940,stringappend_12950) => + (stringappend_12940, stringappend_12950) ) of - (_, stringappend_12590) => - (let stringappend_12600 = (string_drop stringappend_12570 - stringappend_12590) in - (let (imm, stringappend_12620) = + (_, stringappend_12950) => + (let stringappend_12960 = (string_drop stringappend_12930 + stringappend_12950) in + (let (imm, stringappend_12980) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_12600 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_12610,stringappend_12620) => - (stringappend_12610, stringappend_12620) + stringappend_12960 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_12970,stringappend_12980) => + (stringappend_12970, stringappend_12980) )) in - if(((string_drop stringappend_12600 stringappend_12620)) = ('''')) then + if(((string_drop stringappend_12960 stringappend_12980)) = ('''')) then (STORE (imm,rs1,rd,size1,aq,rl)) else undefined)) )))) )))) )))))))) - else if (((((string_startswith stringappend_10760 (''addiw''))) \<and> ((let stringappend_12640 = (string_drop stringappend_10760 ((string_length (''addiw'')))) in - if ((case ((spc_matches_prefix stringappend_12640)) of - Some (stringappend_12650,stringappend_12660) => - (let stringappend_12670 = (string_drop stringappend_12640 stringappend_12660) in - if ((case ((reg_name_matches_prefix stringappend_12670 + else if (((((string_startswith stringappend_11120 (''addiw''))) \<and> ((let stringappend_13000 = (string_drop stringappend_11120 ((string_length (''addiw'')))) in + if ((case ((spc_matches_prefix stringappend_13000)) of + Some (stringappend_13010,stringappend_13020) => + (let stringappend_13030 = (string_drop stringappend_13000 stringappend_13020) in + if ((case ((reg_name_matches_prefix stringappend_13030 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12680,stringappend_12690) => - (let stringappend_12700 = - (string_drop stringappend_12670 stringappend_12690) in - if ((case ((sep_matches_prefix stringappend_12700)) of - Some (stringappend_12710,stringappend_12720) => - (let stringappend_12730 = - (string_drop stringappend_12700 stringappend_12720) in - if ((case ((reg_name_matches_prefix stringappend_12730 + Some (stringappend_13040,stringappend_13050) => + (let stringappend_13060 = + (string_drop stringappend_13030 stringappend_13050) in + if ((case ((sep_matches_prefix stringappend_13060)) of + Some (stringappend_13070,stringappend_13080) => + (let stringappend_13090 = + (string_drop stringappend_13060 stringappend_13080) in + if ((case ((reg_name_matches_prefix stringappend_13090 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12740,stringappend_12750) => - (let stringappend_12760 = - (string_drop stringappend_12730 stringappend_12750) in - if ((case ((sep_matches_prefix stringappend_12760)) of - Some (stringappend_12770,stringappend_12780) => - (let stringappend_12790 = - (string_drop stringappend_12760 stringappend_12780) in + Some (stringappend_13100,stringappend_13110) => + (let stringappend_13120 = + (string_drop stringappend_13090 stringappend_13110) in + if ((case ((sep_matches_prefix stringappend_13120)) of + Some (stringappend_13130,stringappend_13140) => + (let stringappend_13150 = + (string_drop stringappend_13120 stringappend_13140) in if ((case ((hex_bits_12_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_12790 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_13150 :: (( 12 Word.word * ii))option)) of - Some (stringappend_12800,stringappend_12810) => - if(((string_drop stringappend_12790 stringappend_12810)) = ('''')) then + Some (stringappend_13160,stringappend_13170) => + if(((string_drop stringappend_13150 stringappend_13170)) = ('''')) then True else False | None => False )) then @@ -15083,83 +15158,83 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_12640 = (string_drop stringappend_10760 ((string_length (''addiw'')))) in + (let stringappend_13000 = (string_drop stringappend_11120 ((string_length (''addiw'')))) in (case - (case ((spc_matches_prefix stringappend_12640)) of - Some (stringappend_12650,stringappend_12660) => - (stringappend_12650, stringappend_12660) + (case ((spc_matches_prefix stringappend_13000)) of + Some (stringappend_13010,stringappend_13020) => + (stringappend_13010, stringappend_13020) ) of - (_, stringappend_12660) => - (let stringappend_12670 = (string_drop stringappend_12640 - stringappend_12660) in - (let (rd, stringappend_12690) = - ((case ((reg_name_matches_prefix stringappend_12670 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12680,stringappend_12690) => - (stringappend_12680, stringappend_12690) + (_, stringappend_13020) => + (let stringappend_13030 = (string_drop stringappend_13000 + stringappend_13020) in + (let (rd, stringappend_13050) = + ((case ((reg_name_matches_prefix stringappend_13030 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13040,stringappend_13050) => + (stringappend_13040, stringappend_13050) )) in - (let stringappend_12700 = (string_drop stringappend_12670 - stringappend_12690) in + (let stringappend_13060 = (string_drop stringappend_13030 + stringappend_13050) in (case - (case ((sep_matches_prefix stringappend_12700)) of - Some (stringappend_12710,stringappend_12720) => - (stringappend_12710, stringappend_12720) + (case ((sep_matches_prefix stringappend_13060)) of + Some (stringappend_13070,stringappend_13080) => + (stringappend_13070, stringappend_13080) ) of - (_, stringappend_12720) => - (let stringappend_12730 = (string_drop stringappend_12700 - stringappend_12720) in - (let (rs1, stringappend_12750) = - ((case ((reg_name_matches_prefix stringappend_12730 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12740,stringappend_12750) => - (stringappend_12740, stringappend_12750) + (_, stringappend_13080) => + (let stringappend_13090 = (string_drop stringappend_13060 + stringappend_13080) in + (let (rs1, stringappend_13110) = + ((case ((reg_name_matches_prefix stringappend_13090 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13100,stringappend_13110) => + (stringappend_13100, stringappend_13110) )) in - (let stringappend_12760 = (string_drop stringappend_12730 - stringappend_12750) in + (let stringappend_13120 = (string_drop stringappend_13090 + stringappend_13110) in (case - (case ((sep_matches_prefix stringappend_12760)) of - Some (stringappend_12770,stringappend_12780) => - (stringappend_12770, stringappend_12780) + (case ((sep_matches_prefix stringappend_13120)) of + Some (stringappend_13130,stringappend_13140) => + (stringappend_13130, stringappend_13140) ) of - (_, stringappend_12780) => - (let stringappend_12790 = (string_drop stringappend_12760 - stringappend_12780) in - (let (imm, stringappend_12810) = + (_, stringappend_13140) => + (let stringappend_13150 = (string_drop stringappend_13120 + stringappend_13140) in + (let (imm, stringappend_13170) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_12790 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_12800,stringappend_12810) => - (stringappend_12800, stringappend_12810) + stringappend_13150 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_13160,stringappend_13170) => + (stringappend_13160, stringappend_13170) )) in - if(((string_drop stringappend_12790 stringappend_12810)) = ('''')) then + if(((string_drop stringappend_13150 stringappend_13170)) = ('''')) then (ADDIW (imm,rs1,rd)) else undefined)) )))) )))) )) - else if ((case ((shiftw_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_12830,stringappend_12840) => - (let stringappend_12850 = (string_drop stringappend_10760 stringappend_12840) in - if ((case ((spc_matches_prefix stringappend_12850)) of - Some (stringappend_12860,stringappend_12870) => - (let stringappend_12880 = (string_drop stringappend_12850 stringappend_12870) in - if ((case ((reg_name_matches_prefix stringappend_12880 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12890,stringappend_12900) => - (let stringappend_12910 = (string_drop stringappend_12880 stringappend_12900) in - if ((case ((sep_matches_prefix stringappend_12910)) of - Some (stringappend_12920,stringappend_12930) => - (let stringappend_12940 = (string_drop stringappend_12910 stringappend_12930) in - if ((case ((reg_name_matches_prefix stringappend_12940 + else if ((case ((shiftw_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_13190,stringappend_13200) => + (let stringappend_13210 = (string_drop stringappend_11120 stringappend_13200) in + if ((case ((spc_matches_prefix stringappend_13210)) of + Some (stringappend_13220,stringappend_13230) => + (let stringappend_13240 = (string_drop stringappend_13210 stringappend_13230) in + if ((case ((reg_name_matches_prefix stringappend_13240 :: (( 5 Word.word * ii))option)) of + Some (stringappend_13250,stringappend_13260) => + (let stringappend_13270 = (string_drop stringappend_13240 stringappend_13260) in + if ((case ((sep_matches_prefix stringappend_13270)) of + Some (stringappend_13280,stringappend_13290) => + (let stringappend_13300 = (string_drop stringappend_13270 stringappend_13290) in + if ((case ((reg_name_matches_prefix stringappend_13300 :: (( 5 Word.word * ii))option)) of - Some (stringappend_12950,stringappend_12960) => - (let stringappend_12970 = - (string_drop stringappend_12940 stringappend_12960) in - if ((case ((sep_matches_prefix stringappend_12970)) of - Some (stringappend_12980,stringappend_12990) => - (let stringappend_13000 = - (string_drop stringappend_12970 stringappend_12990) in + Some (stringappend_13310,stringappend_13320) => + (let stringappend_13330 = + (string_drop stringappend_13300 stringappend_13320) in + if ((case ((sep_matches_prefix stringappend_13330)) of + Some (stringappend_13340,stringappend_13350) => + (let stringappend_13360 = + (string_drop stringappend_13330 stringappend_13350) in if ((case ((hex_bits_5_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_13000 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_13360 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13010,stringappend_13020) => - if(((string_drop stringappend_13000 stringappend_13020)) = ('''')) then + Some (stringappend_13370,stringappend_13380) => + if(((string_drop stringappend_13360 stringappend_13380)) = ('''')) then True else False | None => False )) then @@ -15187,87 +15262,87 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where else False) | None => False )) then - (let (op1, stringappend_12840) = - ((case ((shiftw_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_12830,stringappend_12840) => - (stringappend_12830, stringappend_12840) + (let (op1, stringappend_13200) = + ((case ((shiftw_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_13190,stringappend_13200) => + (stringappend_13190, stringappend_13200) )) in - (let stringappend_12850 = (string_drop stringappend_10760 stringappend_12840) in + (let stringappend_13210 = (string_drop stringappend_11120 stringappend_13200) in (case - (case ((spc_matches_prefix stringappend_12850)) of - Some (stringappend_12860,stringappend_12870) => - (stringappend_12860, stringappend_12870) + (case ((spc_matches_prefix stringappend_13210)) of + Some (stringappend_13220,stringappend_13230) => + (stringappend_13220, stringappend_13230) ) of - (_, stringappend_12870) => - (let stringappend_12880 = (string_drop stringappend_12850 - stringappend_12870) in - (let (rd, stringappend_12900) = - ((case ((reg_name_matches_prefix stringappend_12880 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12890,stringappend_12900) => - (stringappend_12890, stringappend_12900) + (_, stringappend_13230) => + (let stringappend_13240 = (string_drop stringappend_13210 + stringappend_13230) in + (let (rd, stringappend_13260) = + ((case ((reg_name_matches_prefix stringappend_13240 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13250,stringappend_13260) => + (stringappend_13250, stringappend_13260) )) in - (let stringappend_12910 = (string_drop stringappend_12880 - stringappend_12900) in + (let stringappend_13270 = (string_drop stringappend_13240 + stringappend_13260) in (case - (case ((sep_matches_prefix stringappend_12910)) of - Some (stringappend_12920,stringappend_12930) => - (stringappend_12920, stringappend_12930) + (case ((sep_matches_prefix stringappend_13270)) of + Some (stringappend_13280,stringappend_13290) => + (stringappend_13280, stringappend_13290) ) of - (_, stringappend_12930) => - (let stringappend_12940 = (string_drop stringappend_12910 - stringappend_12930) in - (let (rs1, stringappend_12960) = - ((case ((reg_name_matches_prefix stringappend_12940 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_12950,stringappend_12960) => - (stringappend_12950, stringappend_12960) + (_, stringappend_13290) => + (let stringappend_13300 = (string_drop stringappend_13270 + stringappend_13290) in + (let (rs1, stringappend_13320) = + ((case ((reg_name_matches_prefix stringappend_13300 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13310,stringappend_13320) => + (stringappend_13310, stringappend_13320) )) in - (let stringappend_12970 = (string_drop stringappend_12940 - stringappend_12960) in + (let stringappend_13330 = (string_drop stringappend_13300 + stringappend_13320) in (case - (case ((sep_matches_prefix stringappend_12970)) of - Some (stringappend_12980,stringappend_12990) => - (stringappend_12980, stringappend_12990) + (case ((sep_matches_prefix stringappend_13330)) of + Some (stringappend_13340,stringappend_13350) => + (stringappend_13340, stringappend_13350) ) of - (_, stringappend_12990) => - (let stringappend_13000 = (string_drop stringappend_12970 - stringappend_12990) in - (let (shamt, stringappend_13020) = + (_, stringappend_13350) => + (let stringappend_13360 = (string_drop stringappend_13330 + stringappend_13350) in + (let (shamt, stringappend_13380) = ((case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_13000 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13010,stringappend_13020) => - (stringappend_13010, stringappend_13020) + stringappend_13360 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13370,stringappend_13380) => + (stringappend_13370, stringappend_13380) )) in - if(((string_drop stringappend_13000 stringappend_13020)) = ('''')) then + if(((string_drop stringappend_13360 stringappend_13380)) = ('''')) then (SHIFTW (shamt,rs1,rd,op1)) else undefined)) )))) )))) ))) - else if ((case ((rtypew_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_13040,stringappend_13050) => - (let stringappend_13060 = (string_drop stringappend_10760 stringappend_13050) in - if ((case ((spc_matches_prefix stringappend_13060)) of - Some (stringappend_13070,stringappend_13080) => - (let stringappend_13090 = (string_drop stringappend_13060 stringappend_13080) in - if ((case ((reg_name_matches_prefix stringappend_13090 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13100,stringappend_13110) => - (let stringappend_13120 = (string_drop stringappend_13090 stringappend_13110) in - if ((case ((sep_matches_prefix stringappend_13120)) of - Some (stringappend_13130,stringappend_13140) => - (let stringappend_13150 = (string_drop stringappend_13120 stringappend_13140) in - if ((case ((reg_name_matches_prefix stringappend_13150 + else if ((case ((rtypew_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_13400,stringappend_13410) => + (let stringappend_13420 = (string_drop stringappend_11120 stringappend_13410) in + if ((case ((spc_matches_prefix stringappend_13420)) of + Some (stringappend_13430,stringappend_13440) => + (let stringappend_13450 = (string_drop stringappend_13420 stringappend_13440) in + if ((case ((reg_name_matches_prefix stringappend_13450 :: (( 5 Word.word * ii))option)) of + Some (stringappend_13460,stringappend_13470) => + (let stringappend_13480 = (string_drop stringappend_13450 stringappend_13470) in + if ((case ((sep_matches_prefix stringappend_13480)) of + Some (stringappend_13490,stringappend_13500) => + (let stringappend_13510 = (string_drop stringappend_13480 stringappend_13500) in + if ((case ((reg_name_matches_prefix stringappend_13510 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13160,stringappend_13170) => - (let stringappend_13180 = - (string_drop stringappend_13150 stringappend_13170) in - if ((case ((sep_matches_prefix stringappend_13180)) of - Some (stringappend_13190,stringappend_13200) => - (let stringappend_13210 = - (string_drop stringappend_13180 stringappend_13200) in - if ((case ((reg_name_matches_prefix stringappend_13210 + Some (stringappend_13520,stringappend_13530) => + (let stringappend_13540 = + (string_drop stringappend_13510 stringappend_13530) in + if ((case ((sep_matches_prefix stringappend_13540)) of + Some (stringappend_13550,stringappend_13560) => + (let stringappend_13570 = + (string_drop stringappend_13540 stringappend_13560) in + if ((case ((reg_name_matches_prefix stringappend_13570 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13220,stringappend_13230) => - if(((string_drop stringappend_13210 stringappend_13230)) = ('''')) then + Some (stringappend_13580,stringappend_13590) => + if(((string_drop stringappend_13570 stringappend_13590)) = ('''')) then True else False | None => False )) then @@ -15295,85 +15370,177 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where else False) | None => False )) then - (let (op1, stringappend_13050) = - ((case ((rtypew_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_13040,stringappend_13050) => - (stringappend_13040, stringappend_13050) + (let (op1, stringappend_13410) = + ((case ((rtypew_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_13400,stringappend_13410) => + (stringappend_13400, stringappend_13410) )) in - (let stringappend_13060 = (string_drop stringappend_10760 stringappend_13050) in + (let stringappend_13420 = (string_drop stringappend_11120 stringappend_13410) in (case - (case ((spc_matches_prefix stringappend_13060)) of - Some (stringappend_13070,stringappend_13080) => - (stringappend_13070, stringappend_13080) + (case ((spc_matches_prefix stringappend_13420)) of + Some (stringappend_13430,stringappend_13440) => + (stringappend_13430, stringappend_13440) ) of - (_, stringappend_13080) => - (let stringappend_13090 = (string_drop stringappend_13060 - stringappend_13080) in - (let (rd, stringappend_13110) = - ((case ((reg_name_matches_prefix stringappend_13090 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13100,stringappend_13110) => - (stringappend_13100, stringappend_13110) + (_, stringappend_13440) => + (let stringappend_13450 = (string_drop stringappend_13420 + stringappend_13440) in + (let (rd, stringappend_13470) = + ((case ((reg_name_matches_prefix stringappend_13450 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13460,stringappend_13470) => + (stringappend_13460, stringappend_13470) )) in - (let stringappend_13120 = (string_drop stringappend_13090 - stringappend_13110) in + (let stringappend_13480 = (string_drop stringappend_13450 + stringappend_13470) in (case - (case ((sep_matches_prefix stringappend_13120)) of - Some (stringappend_13130,stringappend_13140) => - (stringappend_13130, stringappend_13140) + (case ((sep_matches_prefix stringappend_13480)) of + Some (stringappend_13490,stringappend_13500) => + (stringappend_13490, stringappend_13500) ) of - (_, stringappend_13140) => - (let stringappend_13150 = (string_drop stringappend_13120 - stringappend_13140) in - (let (rs1, stringappend_13170) = - ((case ((reg_name_matches_prefix stringappend_13150 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13160,stringappend_13170) => - (stringappend_13160, stringappend_13170) + (_, stringappend_13500) => + (let stringappend_13510 = (string_drop stringappend_13480 + stringappend_13500) in + (let (rs1, stringappend_13530) = + ((case ((reg_name_matches_prefix stringappend_13510 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13520,stringappend_13530) => + (stringappend_13520, stringappend_13530) )) in - (let stringappend_13180 = (string_drop stringappend_13150 - stringappend_13170) in + (let stringappend_13540 = (string_drop stringappend_13510 + stringappend_13530) in (case - (case ((sep_matches_prefix stringappend_13180)) of - Some (stringappend_13190,stringappend_13200) => - (stringappend_13190, stringappend_13200) + (case ((sep_matches_prefix stringappend_13540)) of + Some (stringappend_13550,stringappend_13560) => + (stringappend_13550, stringappend_13560) ) of - (_, stringappend_13200) => - (let stringappend_13210 = (string_drop stringappend_13180 - stringappend_13200) in - (let (rs2, stringappend_13230) = - ((case ((reg_name_matches_prefix stringappend_13210 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13220,stringappend_13230) => - (stringappend_13220, stringappend_13230) + (_, stringappend_13560) => + (let stringappend_13570 = (string_drop stringappend_13540 + stringappend_13560) in + (let (rs2, stringappend_13590) = + ((case ((reg_name_matches_prefix stringappend_13570 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13580,stringappend_13590) => + (stringappend_13580, stringappend_13590) )) in - if(((string_drop stringappend_13210 stringappend_13230)) = ('''')) then + if(((string_drop stringappend_13570 stringappend_13590)) = ('''')) then (RTYPEW (rs2,rs1,rd,op1)) else undefined)) )))) )))) ))) - else if ((case ((mul_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_13250,stringappend_13260) => - (let stringappend_13270 = (string_drop stringappend_10760 stringappend_13260) in - if ((case ((spc_matches_prefix stringappend_13270)) of - Some (stringappend_13280,stringappend_13290) => - (let stringappend_13300 = (string_drop stringappend_13270 stringappend_13290) in - if ((case ((reg_name_matches_prefix stringappend_13300 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13310,stringappend_13320) => - (let stringappend_13330 = (string_drop stringappend_13300 stringappend_13320) in - if ((case ((sep_matches_prefix stringappend_13330)) of - Some (stringappend_13340,stringappend_13350) => - (let stringappend_13360 = (string_drop stringappend_13330 stringappend_13350) in - if ((case ((reg_name_matches_prefix stringappend_13360 + else if ((case ((shiftiwop_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_13610,stringappend_13620) => + (let stringappend_13630 = (string_drop stringappend_11120 stringappend_13620) in + if ((case ((spc_matches_prefix stringappend_13630)) of + Some (stringappend_13640,stringappend_13650) => + (let stringappend_13660 = (string_drop stringappend_13630 stringappend_13650) in + if ((case ((reg_name_matches_prefix stringappend_13660 :: (( 5 Word.word * ii))option)) of + Some (stringappend_13670,stringappend_13680) => + (let stringappend_13690 = (string_drop stringappend_13660 stringappend_13680) in + if ((case ((sep_matches_prefix stringappend_13690)) of + Some (stringappend_13700,stringappend_13710) => + (let stringappend_13720 = (string_drop stringappend_13690 stringappend_13710) in + if ((case ((reg_name_matches_prefix stringappend_13720 + :: (( 5 Word.word * ii))option)) of + Some (stringappend_13730,stringappend_13740) => + (let stringappend_13750 = + (string_drop stringappend_13720 stringappend_13740) in + if ((case ((hex_bits_5_matches_prefix + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_13750 + :: (( 5 Word.word * ii))option)) of + Some (stringappend_13760,stringappend_13770) => + if(((string_drop stringappend_13750 stringappend_13770)) = ('''')) then + True else False + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + (let (op1, stringappend_13620) = + ((case ((shiftiwop_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_13610,stringappend_13620) => + (stringappend_13610, stringappend_13620) + )) in + (let stringappend_13630 = (string_drop stringappend_11120 stringappend_13620) in + (case + (case ((spc_matches_prefix stringappend_13630)) of + Some (stringappend_13640,stringappend_13650) => + (stringappend_13640, stringappend_13650) + ) of + (_, stringappend_13650) => + (let stringappend_13660 = (string_drop stringappend_13630 + stringappend_13650) in + (let (rd, stringappend_13680) = + ((case ((reg_name_matches_prefix stringappend_13660 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13670,stringappend_13680) => + (stringappend_13670, stringappend_13680) + )) in + (let stringappend_13690 = (string_drop stringappend_13660 + stringappend_13680) in + (case + (case ((sep_matches_prefix stringappend_13690)) of + Some (stringappend_13700,stringappend_13710) => + (stringappend_13700, stringappend_13710) + ) of + (_, stringappend_13710) => + (let stringappend_13720 = (string_drop stringappend_13690 + stringappend_13710) in + (let (rs1, stringappend_13740) = + ((case ((reg_name_matches_prefix stringappend_13720 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13730,stringappend_13740) => + (stringappend_13730, stringappend_13740) + )) in + (let stringappend_13750 = (string_drop stringappend_13720 + stringappend_13740) in + (let (shamt, stringappend_13770) = + ((case ((hex_bits_5_matches_prefix + instance_Sail2_values_Bitvector_Machine_word_mword_dict + stringappend_13750 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13760,stringappend_13770) => + (stringappend_13760, stringappend_13770) + )) in + if(((string_drop stringappend_13750 stringappend_13770)) = ('''')) then + (SHIFTIWOP (shamt,rs1,rd,op1)) else undefined)))) + )))) + ))) + else if ((case ((mul_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_13790,stringappend_13800) => + (let stringappend_13810 = (string_drop stringappend_11120 stringappend_13800) in + if ((case ((spc_matches_prefix stringappend_13810)) of + Some (stringappend_13820,stringappend_13830) => + (let stringappend_13840 = (string_drop stringappend_13810 stringappend_13830) in + if ((case ((reg_name_matches_prefix stringappend_13840 :: (( 5 Word.word * ii))option)) of + Some (stringappend_13850,stringappend_13860) => + (let stringappend_13870 = (string_drop stringappend_13840 stringappend_13860) in + if ((case ((sep_matches_prefix stringappend_13870)) of + Some (stringappend_13880,stringappend_13890) => + (let stringappend_13900 = (string_drop stringappend_13870 stringappend_13890) in + if ((case ((reg_name_matches_prefix stringappend_13900 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13370,stringappend_13380) => - (let stringappend_13390 = - (string_drop stringappend_13360 stringappend_13380) in - if ((case ((sep_matches_prefix stringappend_13390)) of - Some (stringappend_13400,stringappend_13410) => - (let stringappend_13420 = - (string_drop stringappend_13390 stringappend_13410) in - if ((case ((reg_name_matches_prefix stringappend_13420 + Some (stringappend_13910,stringappend_13920) => + (let stringappend_13930 = + (string_drop stringappend_13900 stringappend_13920) in + if ((case ((sep_matches_prefix stringappend_13930)) of + Some (stringappend_13940,stringappend_13950) => + (let stringappend_13960 = + (string_drop stringappend_13930 stringappend_13950) in + if ((case ((reg_name_matches_prefix stringappend_13960 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13430,stringappend_13440) => - if(((string_drop stringappend_13420 stringappend_13440)) = ('''')) then + Some (stringappend_13970,stringappend_13980) => + if(((string_drop stringappend_13960 stringappend_13980)) = ('''')) then True else False | None => False )) then @@ -15401,90 +15568,90 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where else False) | None => False )) then - (let ((high, signed1, signed2), stringappend_13260) = - ((case ((mul_mnemonic_matches_prefix stringappend_10760)) of - Some (stringappend_13250,stringappend_13260) => - (stringappend_13250, stringappend_13260) + (let ((high, signed1, signed2), stringappend_13800) = + ((case ((mul_mnemonic_matches_prefix stringappend_11120)) of + Some (stringappend_13790,stringappend_13800) => + (stringappend_13790, stringappend_13800) )) in - (let stringappend_13270 = (string_drop stringappend_10760 stringappend_13260) in + (let stringappend_13810 = (string_drop stringappend_11120 stringappend_13800) in (case - (case ((spc_matches_prefix stringappend_13270)) of - Some (stringappend_13280,stringappend_13290) => - (stringappend_13280, stringappend_13290) + (case ((spc_matches_prefix stringappend_13810)) of + Some (stringappend_13820,stringappend_13830) => + (stringappend_13820, stringappend_13830) ) of - (_, stringappend_13290) => - (let stringappend_13300 = (string_drop stringappend_13270 - stringappend_13290) in - (let (rd, stringappend_13320) = - ((case ((reg_name_matches_prefix stringappend_13300 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13310,stringappend_13320) => - (stringappend_13310, stringappend_13320) + (_, stringappend_13830) => + (let stringappend_13840 = (string_drop stringappend_13810 + stringappend_13830) in + (let (rd, stringappend_13860) = + ((case ((reg_name_matches_prefix stringappend_13840 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13850,stringappend_13860) => + (stringappend_13850, stringappend_13860) )) in - (let stringappend_13330 = (string_drop stringappend_13300 - stringappend_13320) in + (let stringappend_13870 = (string_drop stringappend_13840 + stringappend_13860) in (case - (case ((sep_matches_prefix stringappend_13330)) of - Some (stringappend_13340,stringappend_13350) => - (stringappend_13340, stringappend_13350) + (case ((sep_matches_prefix stringappend_13870)) of + Some (stringappend_13880,stringappend_13890) => + (stringappend_13880, stringappend_13890) ) of - (_, stringappend_13350) => - (let stringappend_13360 = (string_drop stringappend_13330 - stringappend_13350) in - (let (rs1, stringappend_13380) = - ((case ((reg_name_matches_prefix stringappend_13360 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13370,stringappend_13380) => - (stringappend_13370, stringappend_13380) + (_, stringappend_13890) => + (let stringappend_13900 = (string_drop stringappend_13870 + stringappend_13890) in + (let (rs1, stringappend_13920) = + ((case ((reg_name_matches_prefix stringappend_13900 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13910,stringappend_13920) => + (stringappend_13910, stringappend_13920) )) in - (let stringappend_13390 = (string_drop stringappend_13360 - stringappend_13380) in + (let stringappend_13930 = (string_drop stringappend_13900 + stringappend_13920) in (case - (case ((sep_matches_prefix stringappend_13390)) of - Some (stringappend_13400,stringappend_13410) => - (stringappend_13400, stringappend_13410) + (case ((sep_matches_prefix stringappend_13930)) of + Some (stringappend_13940,stringappend_13950) => + (stringappend_13940, stringappend_13950) ) of - (_, stringappend_13410) => - (let stringappend_13420 = (string_drop stringappend_13390 - stringappend_13410) in - (let (rs2, stringappend_13440) = - ((case ((reg_name_matches_prefix stringappend_13420 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13430,stringappend_13440) => - (stringappend_13430, stringappend_13440) + (_, stringappend_13950) => + (let stringappend_13960 = (string_drop stringappend_13930 + stringappend_13950) in + (let (rs2, stringappend_13980) = + ((case ((reg_name_matches_prefix stringappend_13960 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_13970,stringappend_13980) => + (stringappend_13970, stringappend_13980) )) in - if(((string_drop stringappend_13420 stringappend_13440)) = ('''')) then + if(((string_drop stringappend_13960 stringappend_13980)) = ('''')) then (MUL (rs2,rs1,rd,high,signed1,signed2)) else undefined)) )))) )))) ))) - else if (((((string_startswith stringappend_10760 (''div''))) \<and> ((let stringappend_13460 = (string_drop stringappend_10760 ((string_length (''div'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_13460)) of - Some (stringappend_13470,stringappend_13480) => - (let stringappend_13490 = (string_drop stringappend_13460 stringappend_13480) in - if ((case ((spc_matches_prefix stringappend_13490)) of - Some (stringappend_13500,stringappend_13510) => - (let stringappend_13520 = - (string_drop stringappend_13490 stringappend_13510) in - if ((case ((reg_name_matches_prefix stringappend_13520 + else if (((((string_startswith stringappend_11120 (''div''))) \<and> ((let stringappend_14000 = (string_drop stringappend_11120 ((string_length (''div'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_14000)) of + Some (stringappend_14010,stringappend_14020) => + (let stringappend_14030 = (string_drop stringappend_14000 stringappend_14020) in + if ((case ((spc_matches_prefix stringappend_14030)) of + Some (stringappend_14040,stringappend_14050) => + (let stringappend_14060 = + (string_drop stringappend_14030 stringappend_14050) in + if ((case ((reg_name_matches_prefix stringappend_14060 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13530,stringappend_13540) => - (let stringappend_13550 = - (string_drop stringappend_13520 stringappend_13540) in - if ((case ((sep_matches_prefix stringappend_13550)) of - Some (stringappend_13560,stringappend_13570) => - (let stringappend_13580 = - (string_drop stringappend_13550 stringappend_13570) in - if ((case ((reg_name_matches_prefix stringappend_13580 + Some (stringappend_14070,stringappend_14080) => + (let stringappend_14090 = + (string_drop stringappend_14060 stringappend_14080) in + if ((case ((sep_matches_prefix stringappend_14090)) of + Some (stringappend_14100,stringappend_14110) => + (let stringappend_14120 = + (string_drop stringappend_14090 stringappend_14110) in + if ((case ((reg_name_matches_prefix stringappend_14120 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13590,stringappend_13600) => - (let stringappend_13610 = - (string_drop stringappend_13580 stringappend_13600) in - if ((case ((sep_matches_prefix stringappend_13610)) of - Some (stringappend_13620,stringappend_13630) => - (let stringappend_13640 = - (string_drop stringappend_13610 stringappend_13630) in - if ((case ((reg_name_matches_prefix stringappend_13640 + Some (stringappend_14130,stringappend_14140) => + (let stringappend_14150 = + (string_drop stringappend_14120 stringappend_14140) in + if ((case ((sep_matches_prefix stringappend_14150)) of + Some (stringappend_14160,stringappend_14170) => + (let stringappend_14180 = + (string_drop stringappend_14150 stringappend_14170) in + if ((case ((reg_name_matches_prefix stringappend_14180 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13650,stringappend_13660) => - if(((string_drop stringappend_13640 stringappend_13660)) = ('''')) then + Some (stringappend_14190,stringappend_14200) => + if(((string_drop stringappend_14180 stringappend_14200)) = ('''')) then True else False | None => False )) then @@ -15514,91 +15681,91 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_13460 = (string_drop stringappend_10760 ((string_length (''div'')))) in - (let (s, stringappend_13480) = - ((case ((maybe_not_u_matches_prefix stringappend_13460)) of - Some (stringappend_13470,stringappend_13480) => - (stringappend_13470, stringappend_13480) + (let stringappend_14000 = (string_drop stringappend_11120 ((string_length (''div'')))) in + (let (s, stringappend_14020) = + ((case ((maybe_not_u_matches_prefix stringappend_14000)) of + Some (stringappend_14010,stringappend_14020) => + (stringappend_14010, stringappend_14020) )) in - (let stringappend_13490 = (string_drop stringappend_13460 stringappend_13480) in + (let stringappend_14030 = (string_drop stringappend_14000 stringappend_14020) in (case - (case ((spc_matches_prefix stringappend_13490)) of - Some (stringappend_13500,stringappend_13510) => - (stringappend_13500, stringappend_13510) + (case ((spc_matches_prefix stringappend_14030)) of + Some (stringappend_14040,stringappend_14050) => + (stringappend_14040, stringappend_14050) ) of - (_, stringappend_13510) => - (let stringappend_13520 = (string_drop stringappend_13490 - stringappend_13510) in - (let (rd, stringappend_13540) = - ((case ((reg_name_matches_prefix stringappend_13520 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13530,stringappend_13540) => - (stringappend_13530, stringappend_13540) + (_, stringappend_14050) => + (let stringappend_14060 = (string_drop stringappend_14030 + stringappend_14050) in + (let (rd, stringappend_14080) = + ((case ((reg_name_matches_prefix stringappend_14060 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14070,stringappend_14080) => + (stringappend_14070, stringappend_14080) )) in - (let stringappend_13550 = (string_drop stringappend_13520 - stringappend_13540) in + (let stringappend_14090 = (string_drop stringappend_14060 + stringappend_14080) in (case - (case ((sep_matches_prefix stringappend_13550)) of - Some (stringappend_13560,stringappend_13570) => - (stringappend_13560, stringappend_13570) + (case ((sep_matches_prefix stringappend_14090)) of + Some (stringappend_14100,stringappend_14110) => + (stringappend_14100, stringappend_14110) ) of - (_, stringappend_13570) => - (let stringappend_13580 = (string_drop stringappend_13550 - stringappend_13570) in - (let (rs1, stringappend_13600) = - ((case ((reg_name_matches_prefix stringappend_13580 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13590,stringappend_13600) => - (stringappend_13590, stringappend_13600) + (_, stringappend_14110) => + (let stringappend_14120 = (string_drop stringappend_14090 + stringappend_14110) in + (let (rs1, stringappend_14140) = + ((case ((reg_name_matches_prefix stringappend_14120 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14130,stringappend_14140) => + (stringappend_14130, stringappend_14140) )) in - (let stringappend_13610 = (string_drop stringappend_13580 - stringappend_13600) in + (let stringappend_14150 = (string_drop stringappend_14120 + stringappend_14140) in (case - (case ((sep_matches_prefix stringappend_13610)) of - Some (stringappend_13620,stringappend_13630) => - (stringappend_13620, stringappend_13630) + (case ((sep_matches_prefix stringappend_14150)) of + Some (stringappend_14160,stringappend_14170) => + (stringappend_14160, stringappend_14170) ) of - (_, stringappend_13630) => - (let stringappend_13640 = (string_drop stringappend_13610 - stringappend_13630) in - (let (rs2, stringappend_13660) = - ((case ((reg_name_matches_prefix stringappend_13640 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13650,stringappend_13660) => - (stringappend_13650, stringappend_13660) + (_, stringappend_14170) => + (let stringappend_14180 = (string_drop stringappend_14150 + stringappend_14170) in + (let (rs2, stringappend_14200) = + ((case ((reg_name_matches_prefix stringappend_14180 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14190,stringappend_14200) => + (stringappend_14190, stringappend_14200) )) in - if(((string_drop stringappend_13640 stringappend_13660)) = ('''')) then + if(((string_drop stringappend_14180 stringappend_14200)) = ('''')) then (DIV (rs2,rs1,rd,s)) else undefined)) )))) )))) )))) - else if (((((string_startswith stringappend_10760 (''rem''))) \<and> ((let stringappend_13680 = (string_drop stringappend_10760 ((string_length (''rem'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_13680)) of - Some (stringappend_13690,stringappend_13700) => - (let stringappend_13710 = (string_drop stringappend_13680 stringappend_13700) in - if ((case ((spc_matches_prefix stringappend_13710)) of - Some (stringappend_13720,stringappend_13730) => - (let stringappend_13740 = - (string_drop stringappend_13710 stringappend_13730) in - if ((case ((reg_name_matches_prefix stringappend_13740 + else if (((((string_startswith stringappend_11120 (''rem''))) \<and> ((let stringappend_14220 = (string_drop stringappend_11120 ((string_length (''rem'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_14220)) of + Some (stringappend_14230,stringappend_14240) => + (let stringappend_14250 = (string_drop stringappend_14220 stringappend_14240) in + if ((case ((spc_matches_prefix stringappend_14250)) of + Some (stringappend_14260,stringappend_14270) => + (let stringappend_14280 = + (string_drop stringappend_14250 stringappend_14270) in + if ((case ((reg_name_matches_prefix stringappend_14280 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13750,stringappend_13760) => - (let stringappend_13770 = - (string_drop stringappend_13740 stringappend_13760) in - if ((case ((sep_matches_prefix stringappend_13770)) of - Some (stringappend_13780,stringappend_13790) => - (let stringappend_13800 = - (string_drop stringappend_13770 stringappend_13790) in - if ((case ((reg_name_matches_prefix stringappend_13800 + Some (stringappend_14290,stringappend_14300) => + (let stringappend_14310 = + (string_drop stringappend_14280 stringappend_14300) in + if ((case ((sep_matches_prefix stringappend_14310)) of + Some (stringappend_14320,stringappend_14330) => + (let stringappend_14340 = + (string_drop stringappend_14310 stringappend_14330) in + if ((case ((reg_name_matches_prefix stringappend_14340 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13810,stringappend_13820) => - (let stringappend_13830 = - (string_drop stringappend_13800 stringappend_13820) in - if ((case ((sep_matches_prefix stringappend_13830)) of - Some (stringappend_13840,stringappend_13850) => - (let stringappend_13860 = - (string_drop stringappend_13830 stringappend_13850) in - if ((case ((reg_name_matches_prefix stringappend_13860 + Some (stringappend_14350,stringappend_14360) => + (let stringappend_14370 = + (string_drop stringappend_14340 stringappend_14360) in + if ((case ((sep_matches_prefix stringappend_14370)) of + Some (stringappend_14380,stringappend_14390) => + (let stringappend_14400 = + (string_drop stringappend_14370 stringappend_14390) in + if ((case ((reg_name_matches_prefix stringappend_14400 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13870,stringappend_13880) => - if(((string_drop stringappend_13860 stringappend_13880)) = ('''')) then + Some (stringappend_14410,stringappend_14420) => + if(((string_drop stringappend_14400 stringappend_14420)) = ('''')) then True else False | None => False )) then @@ -15628,87 +15795,87 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_13680 = (string_drop stringappend_10760 ((string_length (''rem'')))) in - (let (s, stringappend_13700) = - ((case ((maybe_not_u_matches_prefix stringappend_13680)) of - Some (stringappend_13690,stringappend_13700) => - (stringappend_13690, stringappend_13700) + (let stringappend_14220 = (string_drop stringappend_11120 ((string_length (''rem'')))) in + (let (s, stringappend_14240) = + ((case ((maybe_not_u_matches_prefix stringappend_14220)) of + Some (stringappend_14230,stringappend_14240) => + (stringappend_14230, stringappend_14240) )) in - (let stringappend_13710 = (string_drop stringappend_13680 stringappend_13700) in + (let stringappend_14250 = (string_drop stringappend_14220 stringappend_14240) in (case - (case ((spc_matches_prefix stringappend_13710)) of - Some (stringappend_13720,stringappend_13730) => - (stringappend_13720, stringappend_13730) + (case ((spc_matches_prefix stringappend_14250)) of + Some (stringappend_14260,stringappend_14270) => + (stringappend_14260, stringappend_14270) ) of - (_, stringappend_13730) => - (let stringappend_13740 = (string_drop stringappend_13710 - stringappend_13730) in - (let (rd, stringappend_13760) = - ((case ((reg_name_matches_prefix stringappend_13740 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13750,stringappend_13760) => - (stringappend_13750, stringappend_13760) + (_, stringappend_14270) => + (let stringappend_14280 = (string_drop stringappend_14250 + stringappend_14270) in + (let (rd, stringappend_14300) = + ((case ((reg_name_matches_prefix stringappend_14280 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14290,stringappend_14300) => + (stringappend_14290, stringappend_14300) )) in - (let stringappend_13770 = (string_drop stringappend_13740 - stringappend_13760) in + (let stringappend_14310 = (string_drop stringappend_14280 + stringappend_14300) in (case - (case ((sep_matches_prefix stringappend_13770)) of - Some (stringappend_13780,stringappend_13790) => - (stringappend_13780, stringappend_13790) + (case ((sep_matches_prefix stringappend_14310)) of + Some (stringappend_14320,stringappend_14330) => + (stringappend_14320, stringappend_14330) ) of - (_, stringappend_13790) => - (let stringappend_13800 = (string_drop stringappend_13770 - stringappend_13790) in - (let (rs1, stringappend_13820) = - ((case ((reg_name_matches_prefix stringappend_13800 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13810,stringappend_13820) => - (stringappend_13810, stringappend_13820) + (_, stringappend_14330) => + (let stringappend_14340 = (string_drop stringappend_14310 + stringappend_14330) in + (let (rs1, stringappend_14360) = + ((case ((reg_name_matches_prefix stringappend_14340 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14350,stringappend_14360) => + (stringappend_14350, stringappend_14360) )) in - (let stringappend_13830 = (string_drop stringappend_13800 - stringappend_13820) in + (let stringappend_14370 = (string_drop stringappend_14340 + stringappend_14360) in (case - (case ((sep_matches_prefix stringappend_13830)) of - Some (stringappend_13840,stringappend_13850) => - (stringappend_13840, stringappend_13850) + (case ((sep_matches_prefix stringappend_14370)) of + Some (stringappend_14380,stringappend_14390) => + (stringappend_14380, stringappend_14390) ) of - (_, stringappend_13850) => - (let stringappend_13860 = (string_drop stringappend_13830 - stringappend_13850) in - (let (rs2, stringappend_13880) = - ((case ((reg_name_matches_prefix stringappend_13860 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13870,stringappend_13880) => - (stringappend_13870, stringappend_13880) + (_, stringappend_14390) => + (let stringappend_14400 = (string_drop stringappend_14370 + stringappend_14390) in + (let (rs2, stringappend_14420) = + ((case ((reg_name_matches_prefix stringappend_14400 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14410,stringappend_14420) => + (stringappend_14410, stringappend_14420) )) in - if(((string_drop stringappend_13860 stringappend_13880)) = ('''')) then + if(((string_drop stringappend_14400 stringappend_14420)) = ('''')) then (REM (rs2,rs1,rd,s)) else undefined)) )))) )))) )))) - else if (((((string_startswith stringappend_10760 (''mulw''))) \<and> ((let stringappend_13900 = (string_drop stringappend_10760 ((string_length (''mulw'')))) in - if ((case ((spc_matches_prefix stringappend_13900)) of - Some (stringappend_13910,stringappend_13920) => - (let stringappend_13930 = (string_drop stringappend_13900 stringappend_13920) in - if ((case ((reg_name_matches_prefix stringappend_13930 + else if (((((string_startswith stringappend_11120 (''mulw''))) \<and> ((let stringappend_14440 = (string_drop stringappend_11120 ((string_length (''mulw'')))) in + if ((case ((spc_matches_prefix stringappend_14440)) of + Some (stringappend_14450,stringappend_14460) => + (let stringappend_14470 = (string_drop stringappend_14440 stringappend_14460) in + if ((case ((reg_name_matches_prefix stringappend_14470 :: (( 5 Word.word * ii))option)) of - Some (stringappend_13940,stringappend_13950) => - (let stringappend_13960 = - (string_drop stringappend_13930 stringappend_13950) in - if ((case ((sep_matches_prefix stringappend_13960)) of - Some (stringappend_13970,stringappend_13980) => - (let stringappend_13990 = - (string_drop stringappend_13960 stringappend_13980) in - if ((case ((reg_name_matches_prefix stringappend_13990 + Some (stringappend_14480,stringappend_14490) => + (let stringappend_14500 = + (string_drop stringappend_14470 stringappend_14490) in + if ((case ((sep_matches_prefix stringappend_14500)) of + Some (stringappend_14510,stringappend_14520) => + (let stringappend_14530 = + (string_drop stringappend_14500 stringappend_14520) in + if ((case ((reg_name_matches_prefix stringappend_14530 :: (( 5 Word.word * ii))option)) of - Some (stringappend_14000,stringappend_14010) => - (let stringappend_14020 = - (string_drop stringappend_13990 stringappend_14010) in - if ((case ((sep_matches_prefix stringappend_14020)) of - Some (stringappend_14030,stringappend_14040) => - (let stringappend_14050 = - (string_drop stringappend_14020 stringappend_14040) in - if ((case ((reg_name_matches_prefix stringappend_14050 + Some (stringappend_14540,stringappend_14550) => + (let stringappend_14560 = + (string_drop stringappend_14530 stringappend_14550) in + if ((case ((sep_matches_prefix stringappend_14560)) of + Some (stringappend_14570,stringappend_14580) => + (let stringappend_14590 = + (string_drop stringappend_14560 stringappend_14580) in + if ((case ((reg_name_matches_prefix stringappend_14590 :: (( 5 Word.word * ii))option)) of - Some (stringappend_14060,stringappend_14070) => - if(((string_drop stringappend_14050 stringappend_14070)) = ('''')) then + Some (stringappend_14600,stringappend_14610) => + if(((string_drop stringappend_14590 stringappend_14610)) = ('''')) then True else False | None => False )) then @@ -15734,90 +15901,90 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_13900 = (string_drop stringappend_10760 ((string_length (''mulw'')))) in + (let stringappend_14440 = (string_drop stringappend_11120 ((string_length (''mulw'')))) in (case - (case ((spc_matches_prefix stringappend_13900)) of - Some (stringappend_13910,stringappend_13920) => - (stringappend_13910, stringappend_13920) + (case ((spc_matches_prefix stringappend_14440)) of + Some (stringappend_14450,stringappend_14460) => + (stringappend_14450, stringappend_14460) ) of - (_, stringappend_13920) => - (let stringappend_13930 = (string_drop stringappend_13900 - stringappend_13920) in - (let (rd, stringappend_13950) = - ((case ((reg_name_matches_prefix stringappend_13930 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_13940,stringappend_13950) => - (stringappend_13940, stringappend_13950) + (_, stringappend_14460) => + (let stringappend_14470 = (string_drop stringappend_14440 + stringappend_14460) in + (let (rd, stringappend_14490) = + ((case ((reg_name_matches_prefix stringappend_14470 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14480,stringappend_14490) => + (stringappend_14480, stringappend_14490) )) in - (let stringappend_13960 = (string_drop stringappend_13930 - stringappend_13950) in + (let stringappend_14500 = (string_drop stringappend_14470 + stringappend_14490) in (case - (case ((sep_matches_prefix stringappend_13960)) of - Some (stringappend_13970,stringappend_13980) => - (stringappend_13970, stringappend_13980) + (case ((sep_matches_prefix stringappend_14500)) of + Some (stringappend_14510,stringappend_14520) => + (stringappend_14510, stringappend_14520) ) of - (_, stringappend_13980) => - (let stringappend_13990 = (string_drop stringappend_13960 - stringappend_13980) in - (let (rs1, stringappend_14010) = - ((case ((reg_name_matches_prefix stringappend_13990 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14000,stringappend_14010) => - (stringappend_14000, stringappend_14010) + (_, stringappend_14520) => + (let stringappend_14530 = (string_drop stringappend_14500 + stringappend_14520) in + (let (rs1, stringappend_14550) = + ((case ((reg_name_matches_prefix stringappend_14530 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14540,stringappend_14550) => + (stringappend_14540, stringappend_14550) )) in - (let stringappend_14020 = (string_drop stringappend_13990 - stringappend_14010) in + (let stringappend_14560 = (string_drop stringappend_14530 + stringappend_14550) in (case - (case ((sep_matches_prefix stringappend_14020)) of - Some (stringappend_14030,stringappend_14040) => - (stringappend_14030, stringappend_14040) + (case ((sep_matches_prefix stringappend_14560)) of + Some (stringappend_14570,stringappend_14580) => + (stringappend_14570, stringappend_14580) ) of - (_, stringappend_14040) => - (let stringappend_14050 = (string_drop stringappend_14020 - stringappend_14040) in - (let (rs2, stringappend_14070) = - ((case ((reg_name_matches_prefix stringappend_14050 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14060,stringappend_14070) => - (stringappend_14060, stringappend_14070) + (_, stringappend_14580) => + (let stringappend_14590 = (string_drop stringappend_14560 + stringappend_14580) in + (let (rs2, stringappend_14610) = + ((case ((reg_name_matches_prefix stringappend_14590 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14600,stringappend_14610) => + (stringappend_14600, stringappend_14610) )) in - if(((string_drop stringappend_14050 stringappend_14070)) = ('''')) then + if(((string_drop stringappend_14590 stringappend_14610)) = ('''')) then (MULW (rs2,rs1,rd)) else undefined)) )))) )))) )) - else if (((((string_startswith stringappend_10760 (''div''))) \<and> ((let stringappend_14090 = (string_drop stringappend_10760 ((string_length (''div'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_14090)) of - Some (stringappend_14100,stringappend_14110) => - (let stringappend_14120 = (string_drop stringappend_14090 stringappend_14110) in - if (((((string_startswith stringappend_14120 (''w''))) \<and> ((let stringappend_14130 = - (string_drop stringappend_14120 ((string_length (''w'')))) in - if ((case ((spc_matches_prefix stringappend_14130)) of - Some (stringappend_14140,stringappend_14150) => - (let stringappend_14160 = - (string_drop stringappend_14130 stringappend_14150) in - if ((case ((reg_name_matches_prefix stringappend_14160 + else if (((((string_startswith stringappend_11120 (''div''))) \<and> ((let stringappend_14630 = (string_drop stringappend_11120 ((string_length (''div'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_14630)) of + Some (stringappend_14640,stringappend_14650) => + (let stringappend_14660 = (string_drop stringappend_14630 stringappend_14650) in + if (((((string_startswith stringappend_14660 (''w''))) \<and> ((let stringappend_14670 = + (string_drop stringappend_14660 ((string_length (''w'')))) in + if ((case ((spc_matches_prefix stringappend_14670)) of + Some (stringappend_14680,stringappend_14690) => + (let stringappend_14700 = + (string_drop stringappend_14670 stringappend_14690) in + if ((case ((reg_name_matches_prefix stringappend_14700 :: (( 5 Word.word * ii))option)) of - Some (stringappend_14170,stringappend_14180) => - (let stringappend_14190 = - (string_drop stringappend_14160 stringappend_14180) in - if ((case ((sep_matches_prefix stringappend_14190)) of - Some (stringappend_14200,stringappend_14210) => - (let stringappend_14220 = - (string_drop stringappend_14190 stringappend_14210) in - if ((case ((reg_name_matches_prefix stringappend_14220 + Some (stringappend_14710,stringappend_14720) => + (let stringappend_14730 = + (string_drop stringappend_14700 stringappend_14720) in + if ((case ((sep_matches_prefix stringappend_14730)) of + Some (stringappend_14740,stringappend_14750) => + (let stringappend_14760 = + (string_drop stringappend_14730 stringappend_14750) in + if ((case ((reg_name_matches_prefix stringappend_14760 :: (( 5 Word.word * ii))option)) of - Some (stringappend_14230,stringappend_14240) => - (let stringappend_14250 = - (string_drop stringappend_14220 stringappend_14240) in - if ((case ((sep_matches_prefix stringappend_14250)) of - Some (stringappend_14260,stringappend_14270) => - (let stringappend_14280 = - (string_drop stringappend_14250 - stringappend_14270) in + Some (stringappend_14770,stringappend_14780) => + (let stringappend_14790 = + (string_drop stringappend_14760 stringappend_14780) in + if ((case ((sep_matches_prefix stringappend_14790)) of + Some (stringappend_14800,stringappend_14810) => + (let stringappend_14820 = + (string_drop stringappend_14790 + stringappend_14810) in if ((case ((reg_name_matches_prefix - stringappend_14280 + stringappend_14820 :: (( 5 Word.word * ii))option)) of Some - (stringappend_14290,stringappend_14300) => - if(((string_drop stringappend_14280 stringappend_14300)) = ('''')) then + (stringappend_14830,stringappend_14840) => + if(((string_drop stringappend_14820 stringappend_14840)) = ('''')) then True else False | None => False )) then @@ -15849,97 +16016,97 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_14090 = (string_drop stringappend_10760 ((string_length (''div'')))) in - (let (s, stringappend_14110) = - ((case ((maybe_not_u_matches_prefix stringappend_14090)) of - Some (stringappend_14100,stringappend_14110) => - (stringappend_14100, stringappend_14110) + (let stringappend_14630 = (string_drop stringappend_11120 ((string_length (''div'')))) in + (let (s, stringappend_14650) = + ((case ((maybe_not_u_matches_prefix stringappend_14630)) of + Some (stringappend_14640,stringappend_14650) => + (stringappend_14640, stringappend_14650) )) in - (let stringappend_14120 = (string_drop stringappend_14090 stringappend_14110) in - (let stringappend_14130 = (string_drop stringappend_14120 ((string_length (''w'')))) in + (let stringappend_14660 = (string_drop stringappend_14630 stringappend_14650) in + (let stringappend_14670 = (string_drop stringappend_14660 ((string_length (''w'')))) in (case - (case ((spc_matches_prefix stringappend_14130)) of - Some (stringappend_14140,stringappend_14150) => - (stringappend_14140, stringappend_14150) + (case ((spc_matches_prefix stringappend_14670)) of + Some (stringappend_14680,stringappend_14690) => + (stringappend_14680, stringappend_14690) ) of - (_, stringappend_14150) => - (let stringappend_14160 = (string_drop stringappend_14130 - stringappend_14150) in - (let (rd, stringappend_14180) = - ((case ((reg_name_matches_prefix stringappend_14160 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14170,stringappend_14180) => - (stringappend_14170, stringappend_14180) + (_, stringappend_14690) => + (let stringappend_14700 = (string_drop stringappend_14670 + stringappend_14690) in + (let (rd, stringappend_14720) = + ((case ((reg_name_matches_prefix stringappend_14700 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14710,stringappend_14720) => + (stringappend_14710, stringappend_14720) )) in - (let stringappend_14190 = (string_drop stringappend_14160 - stringappend_14180) in + (let stringappend_14730 = (string_drop stringappend_14700 + stringappend_14720) in (case - (case ((sep_matches_prefix stringappend_14190)) of - Some (stringappend_14200,stringappend_14210) => - (stringappend_14200, stringappend_14210) + (case ((sep_matches_prefix stringappend_14730)) of + Some (stringappend_14740,stringappend_14750) => + (stringappend_14740, stringappend_14750) ) of - (_, stringappend_14210) => - (let stringappend_14220 = (string_drop stringappend_14190 - stringappend_14210) in - (let (rs1, stringappend_14240) = - ((case ((reg_name_matches_prefix stringappend_14220 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14230,stringappend_14240) => - (stringappend_14230, stringappend_14240) + (_, stringappend_14750) => + (let stringappend_14760 = (string_drop stringappend_14730 + stringappend_14750) in + (let (rs1, stringappend_14780) = + ((case ((reg_name_matches_prefix stringappend_14760 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14770,stringappend_14780) => + (stringappend_14770, stringappend_14780) )) in - (let stringappend_14250 = (string_drop stringappend_14220 - stringappend_14240) in + (let stringappend_14790 = (string_drop stringappend_14760 + stringappend_14780) in (case - (case ((sep_matches_prefix stringappend_14250)) of - Some (stringappend_14260,stringappend_14270) => - (stringappend_14260, stringappend_14270) + (case ((sep_matches_prefix stringappend_14790)) of + Some (stringappend_14800,stringappend_14810) => + (stringappend_14800, stringappend_14810) ) of - (_, stringappend_14270) => - (let stringappend_14280 = (string_drop stringappend_14250 - stringappend_14270) in - (let (rs2, stringappend_14300) = - ((case ((reg_name_matches_prefix stringappend_14280 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14290,stringappend_14300) => - (stringappend_14290, stringappend_14300) + (_, stringappend_14810) => + (let stringappend_14820 = (string_drop stringappend_14790 + stringappend_14810) in + (let (rs2, stringappend_14840) = + ((case ((reg_name_matches_prefix stringappend_14820 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14830,stringappend_14840) => + (stringappend_14830, stringappend_14840) )) in - if(((string_drop stringappend_14280 stringappend_14300)) = ('''')) then + if(((string_drop stringappend_14820 stringappend_14840)) = ('''')) then (DIVW (rs2,rs1,rd,s)) else undefined)) )))) )))) ))))) - else if (((((string_startswith stringappend_10760 (''rem''))) \<and> ((let stringappend_14320 = (string_drop stringappend_10760 ((string_length (''rem'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_14320)) of - Some (stringappend_14330,stringappend_14340) => - (let stringappend_14350 = (string_drop stringappend_14320 stringappend_14340) in - if (((((string_startswith stringappend_14350 (''w''))) \<and> ((let stringappend_14360 = - (string_drop stringappend_14350 ((string_length (''w'')))) in - if ((case ((spc_matches_prefix stringappend_14360)) of - Some (stringappend_14370,stringappend_14380) => - (let stringappend_14390 = - (string_drop stringappend_14360 stringappend_14380) in - if ((case ((reg_name_matches_prefix stringappend_14390 + else if (((((string_startswith stringappend_11120 (''rem''))) \<and> ((let stringappend_14860 = (string_drop stringappend_11120 ((string_length (''rem'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_14860)) of + Some (stringappend_14870,stringappend_14880) => + (let stringappend_14890 = (string_drop stringappend_14860 stringappend_14880) in + if (((((string_startswith stringappend_14890 (''w''))) \<and> ((let stringappend_14900 = + (string_drop stringappend_14890 ((string_length (''w'')))) in + if ((case ((spc_matches_prefix stringappend_14900)) of + Some (stringappend_14910,stringappend_14920) => + (let stringappend_14930 = + (string_drop stringappend_14900 stringappend_14920) in + if ((case ((reg_name_matches_prefix stringappend_14930 :: (( 5 Word.word * ii))option)) of - Some (stringappend_14400,stringappend_14410) => - (let stringappend_14420 = - (string_drop stringappend_14390 stringappend_14410) in - if ((case ((sep_matches_prefix stringappend_14420)) of - Some (stringappend_14430,stringappend_14440) => - (let stringappend_14450 = - (string_drop stringappend_14420 stringappend_14440) in - if ((case ((reg_name_matches_prefix stringappend_14450 + Some (stringappend_14940,stringappend_14950) => + (let stringappend_14960 = + (string_drop stringappend_14930 stringappend_14950) in + if ((case ((sep_matches_prefix stringappend_14960)) of + Some (stringappend_14970,stringappend_14980) => + (let stringappend_14990 = + (string_drop stringappend_14960 stringappend_14980) in + if ((case ((reg_name_matches_prefix stringappend_14990 :: (( 5 Word.word * ii))option)) of - Some (stringappend_14460,stringappend_14470) => - (let stringappend_14480 = - (string_drop stringappend_14450 stringappend_14470) in - if ((case ((sep_matches_prefix stringappend_14480)) of - Some (stringappend_14490,stringappend_14500) => - (let stringappend_14510 = - (string_drop stringappend_14480 - stringappend_14500) in + Some (stringappend_15000,stringappend_15010) => + (let stringappend_15020 = + (string_drop stringappend_14990 stringappend_15010) in + if ((case ((sep_matches_prefix stringappend_15020)) of + Some (stringappend_15030,stringappend_15040) => + (let stringappend_15050 = + (string_drop stringappend_15020 + stringappend_15040) in if ((case ((reg_name_matches_prefix - stringappend_14510 + stringappend_15050 :: (( 5 Word.word * ii))option)) of Some - (stringappend_14520,stringappend_14530) => - if(((string_drop stringappend_14510 stringappend_14530)) = ('''')) then + (stringappend_15060,stringappend_15070) => + if(((string_drop stringappend_15050 stringappend_15070)) = ('''')) then True else False | None => False )) then @@ -15971,79 +16138,79 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_14320 = (string_drop stringappend_10760 ((string_length (''rem'')))) in - (let (s, stringappend_14340) = - ((case ((maybe_not_u_matches_prefix stringappend_14320)) of - Some (stringappend_14330,stringappend_14340) => - (stringappend_14330, stringappend_14340) + (let stringappend_14860 = (string_drop stringappend_11120 ((string_length (''rem'')))) in + (let (s, stringappend_14880) = + ((case ((maybe_not_u_matches_prefix stringappend_14860)) of + Some (stringappend_14870,stringappend_14880) => + (stringappend_14870, stringappend_14880) )) in - (let stringappend_14350 = (string_drop stringappend_14320 stringappend_14340) in - (let stringappend_14360 = (string_drop stringappend_14350 ((string_length (''w'')))) in + (let stringappend_14890 = (string_drop stringappend_14860 stringappend_14880) in + (let stringappend_14900 = (string_drop stringappend_14890 ((string_length (''w'')))) in (case - (case ((spc_matches_prefix stringappend_14360)) of - Some (stringappend_14370,stringappend_14380) => - (stringappend_14370, stringappend_14380) + (case ((spc_matches_prefix stringappend_14900)) of + Some (stringappend_14910,stringappend_14920) => + (stringappend_14910, stringappend_14920) ) of - (_, stringappend_14380) => - (let stringappend_14390 = (string_drop stringappend_14360 - stringappend_14380) in - (let (rd, stringappend_14410) = - ((case ((reg_name_matches_prefix stringappend_14390 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14400,stringappend_14410) => - (stringappend_14400, stringappend_14410) + (_, stringappend_14920) => + (let stringappend_14930 = (string_drop stringappend_14900 + stringappend_14920) in + (let (rd, stringappend_14950) = + ((case ((reg_name_matches_prefix stringappend_14930 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_14940,stringappend_14950) => + (stringappend_14940, stringappend_14950) )) in - (let stringappend_14420 = (string_drop stringappend_14390 - stringappend_14410) in + (let stringappend_14960 = (string_drop stringappend_14930 + stringappend_14950) in (case - (case ((sep_matches_prefix stringappend_14420)) of - Some (stringappend_14430,stringappend_14440) => - (stringappend_14430, stringappend_14440) + (case ((sep_matches_prefix stringappend_14960)) of + Some (stringappend_14970,stringappend_14980) => + (stringappend_14970, stringappend_14980) ) of - (_, stringappend_14440) => - (let stringappend_14450 = (string_drop stringappend_14420 - stringappend_14440) in - (let (rs1, stringappend_14470) = - ((case ((reg_name_matches_prefix stringappend_14450 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14460,stringappend_14470) => - (stringappend_14460, stringappend_14470) + (_, stringappend_14980) => + (let stringappend_14990 = (string_drop stringappend_14960 + stringappend_14980) in + (let (rs1, stringappend_15010) = + ((case ((reg_name_matches_prefix stringappend_14990 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15000,stringappend_15010) => + (stringappend_15000, stringappend_15010) )) in - (let stringappend_14480 = (string_drop stringappend_14450 - stringappend_14470) in + (let stringappend_15020 = (string_drop stringappend_14990 + stringappend_15010) in (case - (case ((sep_matches_prefix stringappend_14480)) of - Some (stringappend_14490,stringappend_14500) => - (stringappend_14490, stringappend_14500) + (case ((sep_matches_prefix stringappend_15020)) of + Some (stringappend_15030,stringappend_15040) => + (stringappend_15030, stringappend_15040) ) of - (_, stringappend_14500) => - (let stringappend_14510 = (string_drop stringappend_14480 - stringappend_14500) in - (let (rs2, stringappend_14530) = - ((case ((reg_name_matches_prefix stringappend_14510 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14520,stringappend_14530) => - (stringappend_14520, stringappend_14530) + (_, stringappend_15040) => + (let stringappend_15050 = (string_drop stringappend_15020 + stringappend_15040) in + (let (rs2, stringappend_15070) = + ((case ((reg_name_matches_prefix stringappend_15050 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15060,stringappend_15070) => + (stringappend_15060, stringappend_15070) )) in - if(((string_drop stringappend_14510 stringappend_14530)) = ('''')) then + if(((string_drop stringappend_15050 stringappend_15070)) = ('''')) then (REMW (rs2,rs1,rd,s)) else undefined)) )))) )))) ))))) - else if (((((string_startswith stringappend_10760 (''fence''))) \<and> ((let stringappend_14550 = (string_drop stringappend_10760 ((string_length (''fence'')))) in - if ((case ((spc_matches_prefix stringappend_14550)) of - Some (stringappend_14560,stringappend_14570) => - (let stringappend_14580 = (string_drop stringappend_14550 stringappend_14570) in - if ((case ((fence_bits_matches_prefix stringappend_14580 + else if (((((string_startswith stringappend_11120 (''fence''))) \<and> ((let stringappend_15090 = (string_drop stringappend_11120 ((string_length (''fence'')))) in + if ((case ((spc_matches_prefix stringappend_15090)) of + Some (stringappend_15100,stringappend_15110) => + (let stringappend_15120 = (string_drop stringappend_15090 stringappend_15110) in + if ((case ((fence_bits_matches_prefix stringappend_15120 :: (( 4 Word.word * ii))option)) of - Some (stringappend_14590,stringappend_14600) => - (let stringappend_14610 = - (string_drop stringappend_14580 stringappend_14600) in - if ((case ((sep_matches_prefix stringappend_14610)) of - Some (stringappend_14620,stringappend_14630) => - (let stringappend_14640 = - (string_drop stringappend_14610 stringappend_14630) in - if ((case ((fence_bits_matches_prefix stringappend_14640 + Some (stringappend_15130,stringappend_15140) => + (let stringappend_15150 = + (string_drop stringappend_15120 stringappend_15140) in + if ((case ((sep_matches_prefix stringappend_15150)) of + Some (stringappend_15160,stringappend_15170) => + (let stringappend_15180 = + (string_drop stringappend_15150 stringappend_15170) in + if ((case ((fence_bits_matches_prefix stringappend_15180 :: (( 4 Word.word * ii))option)) of - Some (stringappend_14650,stringappend_14660) => - if(((string_drop stringappend_14640 stringappend_14660)) = ('''')) then + Some (stringappend_15190,stringappend_15200) => + if(((string_drop stringappend_15180 stringappend_15200)) = ('''')) then True else False | None => False )) then @@ -16061,80 +16228,80 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False))))) then - (let stringappend_14550 = (string_drop stringappend_10760 ((string_length (''fence'')))) in + (let stringappend_15090 = (string_drop stringappend_11120 ((string_length (''fence'')))) in (case - (case ((spc_matches_prefix stringappend_14550)) of - Some (stringappend_14560,stringappend_14570) => - (stringappend_14560, stringappend_14570) + (case ((spc_matches_prefix stringappend_15090)) of + Some (stringappend_15100,stringappend_15110) => + (stringappend_15100, stringappend_15110) ) of - (_, stringappend_14570) => - (let stringappend_14580 = (string_drop stringappend_14550 - stringappend_14570) in - (let (pred, stringappend_14600) = - ((case ((fence_bits_matches_prefix stringappend_14580 :: (( 4 Word.word * ii)) option)) of - Some (stringappend_14590,stringappend_14600) => - (stringappend_14590, stringappend_14600) + (_, stringappend_15110) => + (let stringappend_15120 = (string_drop stringappend_15090 + stringappend_15110) in + (let (pred, stringappend_15140) = + ((case ((fence_bits_matches_prefix stringappend_15120 :: (( 4 Word.word * ii)) option)) of + Some (stringappend_15130,stringappend_15140) => + (stringappend_15130, stringappend_15140) )) in - (let stringappend_14610 = (string_drop stringappend_14580 - stringappend_14600) in + (let stringappend_15150 = (string_drop stringappend_15120 + stringappend_15140) in (case - (case ((sep_matches_prefix stringappend_14610)) of - Some (stringappend_14620,stringappend_14630) => - (stringappend_14620, stringappend_14630) + (case ((sep_matches_prefix stringappend_15150)) of + Some (stringappend_15160,stringappend_15170) => + (stringappend_15160, stringappend_15170) ) of - (_, stringappend_14630) => - (let stringappend_14640 = (string_drop stringappend_14610 - stringappend_14630) in - (let (succ, stringappend_14660) = - ((case ((fence_bits_matches_prefix stringappend_14640 :: (( 4 Word.word * ii)) option)) of - Some (stringappend_14650,stringappend_14660) => - (stringappend_14650, stringappend_14660) + (_, stringappend_15170) => + (let stringappend_15180 = (string_drop stringappend_15150 + stringappend_15170) in + (let (succ, stringappend_15200) = + ((case ((fence_bits_matches_prefix stringappend_15180 :: (( 4 Word.word * ii)) option)) of + Some (stringappend_15190,stringappend_15200) => + (stringappend_15190, stringappend_15200) )) in - if(((string_drop stringappend_14640 stringappend_14660)) = ('''')) then + if(((string_drop stringappend_15180 stringappend_15200)) = ('''')) then (FENCE (pred,succ)) else undefined)) )))) )) - else if(stringappend_10760 = (''fence.i'')) then (FENCEI () ) else + else if(stringappend_11120 = (''fence.i'')) then (FENCEI () ) else ( - if(stringappend_10760 = (''ecall'')) then (ECALL () ) else + if(stringappend_11120 = (''ecall'')) then (ECALL () ) else ( - if(stringappend_10760 = (''mret'')) then (MRET () ) else + if(stringappend_11120 = (''mret'')) then (MRET () ) else ( - if(stringappend_10760 = (''sret'')) then (SRET () ) else + if(stringappend_11120 = (''sret'')) then (SRET () ) else ( - if(stringappend_10760 = (''ebreak'')) then (EBREAK () ) else + if(stringappend_11120 = (''ebreak'')) then (EBREAK () ) else ( - if(stringappend_10760 = (''wfi'')) then (WFI () ) else + if(stringappend_11120 = (''wfi'')) then (WFI () ) else ( - if (((((string_startswith stringappend_10760 (''sfence.vma''))) + if (((((string_startswith stringappend_11120 (''sfence.vma''))) \<and> - ((let stringappend_14680 = - (string_drop stringappend_10760 + ((let stringappend_15220 = + (string_drop stringappend_11120 ((string_length (''sfence.vma'')))) in - if ((case ((spc_matches_prefix stringappend_14680)) of - Some (stringappend_14690,stringappend_14700) => - (let stringappend_14710 = (string_drop - stringappend_14680 - stringappend_14700) in + if ((case ((spc_matches_prefix stringappend_15220)) of + Some (stringappend_15230,stringappend_15240) => + (let stringappend_15250 = (string_drop + stringappend_15220 + stringappend_15240) in if ((case ((reg_name_matches_prefix - stringappend_14710 + stringappend_15250 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14720,stringappend_14730) => - (let stringappend_14740 = - (string_drop stringappend_14710 - stringappend_14730) in + Some (stringappend_15260,stringappend_15270) => + (let stringappend_15280 = + (string_drop stringappend_15250 + stringappend_15270) in if ((case ((sep_matches_prefix - stringappend_14740)) of - Some (stringappend_14750,stringappend_14760) => - (let stringappend_14770 = - (string_drop stringappend_14740 - stringappend_14760) in + stringappend_15280)) of + Some (stringappend_15290,stringappend_15300) => + (let stringappend_15310 = + (string_drop stringappend_15280 + stringappend_15300) in if ((case ((reg_name_matches_prefix - stringappend_14770 + stringappend_15310 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14780,stringappend_14790) => - if(((string_drop stringappend_14770 - stringappend_14790)) = ('''')) then + Some (stringappend_15320,stringappend_15330) => + if(((string_drop stringappend_15310 + stringappend_15330)) = ('''')) then True else False | None => False )) then True else False) @@ -16144,91 +16311,91 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False) | None => False )) then True else False))))) then - (let stringappend_14680 = (string_drop stringappend_10760 + (let stringappend_15220 = (string_drop stringappend_11120 ((string_length (''sfence.vma'')))) in (case - (case ((spc_matches_prefix stringappend_14680)) of - Some (stringappend_14690,stringappend_14700) => - (stringappend_14690, stringappend_14700) + (case ((spc_matches_prefix stringappend_15220)) of + Some (stringappend_15230,stringappend_15240) => + (stringappend_15230, stringappend_15240) ) of - (_, stringappend_14700) => - (let stringappend_14710 = (string_drop stringappend_14680 - stringappend_14700) in - (let (rs1, stringappend_14730) = - ((case ((reg_name_matches_prefix stringappend_14710 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14720,stringappend_14730) => - (stringappend_14720, stringappend_14730) + (_, stringappend_15240) => + (let stringappend_15250 = (string_drop stringappend_15220 + stringappend_15240) in + (let (rs1, stringappend_15270) = + ((case ((reg_name_matches_prefix stringappend_15250 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15260,stringappend_15270) => + (stringappend_15260, stringappend_15270) )) in - (let stringappend_14740 = (string_drop stringappend_14710 - stringappend_14730) in + (let stringappend_15280 = (string_drop stringappend_15250 + stringappend_15270) in (case - (case ((sep_matches_prefix stringappend_14740)) of - Some (stringappend_14750,stringappend_14760) => - (stringappend_14750, stringappend_14760) + (case ((sep_matches_prefix stringappend_15280)) of + Some (stringappend_15290,stringappend_15300) => + (stringappend_15290, stringappend_15300) ) of - (_, stringappend_14760) => - (let stringappend_14770 = (string_drop stringappend_14740 - stringappend_14760) in - (let (rs2, stringappend_14790) = - ((case ((reg_name_matches_prefix stringappend_14770 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14780,stringappend_14790) => - (stringappend_14780, stringappend_14790) + (_, stringappend_15300) => + (let stringappend_15310 = (string_drop stringappend_15280 + stringappend_15300) in + (let (rs2, stringappend_15330) = + ((case ((reg_name_matches_prefix stringappend_15310 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15320,stringappend_15330) => + (stringappend_15320, stringappend_15330) )) in - if(((string_drop stringappend_14770 stringappend_14790)) = + if(((string_drop stringappend_15310 stringappend_15330)) = ('''')) then (SFENCE_VMA (rs1,rs2)) else undefined)) )))) )) else - if (((((string_startswith stringappend_10760 (''lr.''))) + if (((((string_startswith stringappend_11120 (''lr.''))) \<and> - ((let stringappend_14810 = - (string_drop stringappend_10760 + ((let stringappend_15350 = + (string_drop stringappend_11120 ((string_length (''lr.'')))) in if ((case ((maybe_aq_matches_prefix - stringappend_14810)) of - Some (stringappend_14820,stringappend_14830) => - (let stringappend_14840 = - (string_drop stringappend_14810 - stringappend_14830) in + stringappend_15350)) of + Some (stringappend_15360,stringappend_15370) => + (let stringappend_15380 = + (string_drop stringappend_15350 + stringappend_15370) in if ((case ((maybe_rl_matches_prefix - stringappend_14840)) of - Some (stringappend_14850,stringappend_14860) => - (let stringappend_14870 = - (string_drop stringappend_14840 - stringappend_14860) in + stringappend_15380)) of + Some (stringappend_15390,stringappend_15400) => + (let stringappend_15410 = + (string_drop stringappend_15380 + stringappend_15400) in if ((case ((size_mnemonic_matches_prefix - stringappend_14870)) of - Some (stringappend_14880,stringappend_14890) => - (let stringappend_14900 = - (string_drop stringappend_14870 - stringappend_14890) in + stringappend_15410)) of + Some (stringappend_15420,stringappend_15430) => + (let stringappend_15440 = + (string_drop stringappend_15410 + stringappend_15430) in if ((case ((spc_matches_prefix - stringappend_14900)) of - Some (stringappend_14910,stringappend_14920) => - (let stringappend_14930 = - (string_drop stringappend_14900 - stringappend_14920) in + stringappend_15440)) of + Some (stringappend_15450,stringappend_15460) => + (let stringappend_15470 = + (string_drop stringappend_15440 + stringappend_15460) in if ((case ((reg_name_matches_prefix - stringappend_14930 + stringappend_15470 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14940,stringappend_14950) => - (let stringappend_14960 = + Some (stringappend_15480,stringappend_15490) => + (let stringappend_15500 = (string_drop - stringappend_14930 - stringappend_14950) in + stringappend_15470 + stringappend_15490) in if ((case ((sep_matches_prefix - stringappend_14960)) of - Some (stringappend_14970,stringappend_14980) => - (let stringappend_14990 = + stringappend_15500)) of + Some (stringappend_15510,stringappend_15520) => + (let stringappend_15530 = (string_drop - stringappend_14960 - stringappend_14980) in + stringappend_15500 + stringappend_15520) in if ((case ((reg_name_matches_prefix - stringappend_14990 + stringappend_15530 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15000,stringappend_15010) => + Some (stringappend_15540,stringappend_15550) => if(((string_drop - stringappend_14990 - stringappend_15010)) + stringappend_15530 + stringappend_15550)) = ('''')) then True else False | None => False @@ -16246,133 +16413,133 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False) | None => False )) then True else False))))) then - (let stringappend_14810 = (string_drop stringappend_10760 + (let stringappend_15350 = (string_drop stringappend_11120 ((string_length (''lr.'')))) in - (let (aq, stringappend_14830) = - ((case ((maybe_aq_matches_prefix stringappend_14810)) of - Some (stringappend_14820,stringappend_14830) => - (stringappend_14820, stringappend_14830) + (let (aq, stringappend_15370) = + ((case ((maybe_aq_matches_prefix stringappend_15350)) of + Some (stringappend_15360,stringappend_15370) => + (stringappend_15360, stringappend_15370) )) in - (let stringappend_14840 = (string_drop stringappend_14810 - stringappend_14830) in - (let (rl, stringappend_14860) = - ((case ((maybe_rl_matches_prefix stringappend_14840)) of - Some (stringappend_14850,stringappend_14860) => - (stringappend_14850, stringappend_14860) + (let stringappend_15380 = (string_drop stringappend_15350 + stringappend_15370) in + (let (rl, stringappend_15400) = + ((case ((maybe_rl_matches_prefix stringappend_15380)) of + Some (stringappend_15390,stringappend_15400) => + (stringappend_15390, stringappend_15400) )) in - (let stringappend_14870 = (string_drop stringappend_14840 - stringappend_14860) in - (let (size1, stringappend_14890) = + (let stringappend_15410 = (string_drop stringappend_15380 + stringappend_15400) in + (let (size1, stringappend_15430) = ((case ((size_mnemonic_matches_prefix - stringappend_14870)) of - Some (stringappend_14880,stringappend_14890) => - (stringappend_14880, stringappend_14890) + stringappend_15410)) of + Some (stringappend_15420,stringappend_15430) => + (stringappend_15420, stringappend_15430) )) in - (let stringappend_14900 = (string_drop stringappend_14870 - stringappend_14890) in + (let stringappend_15440 = (string_drop stringappend_15410 + stringappend_15430) in (case - (case ((spc_matches_prefix stringappend_14900)) of - Some (stringappend_14910,stringappend_14920) => - (stringappend_14910, stringappend_14920) + (case ((spc_matches_prefix stringappend_15440)) of + Some (stringappend_15450,stringappend_15460) => + (stringappend_15450, stringappend_15460) ) of - (_, stringappend_14920) => - (let stringappend_14930 = (string_drop stringappend_14900 - stringappend_14920) in - (let (rd, stringappend_14950) = - ((case ((reg_name_matches_prefix stringappend_14930 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_14940,stringappend_14950) => - (stringappend_14940, stringappend_14950) + (_, stringappend_15460) => + (let stringappend_15470 = (string_drop stringappend_15440 + stringappend_15460) in + (let (rd, stringappend_15490) = + ((case ((reg_name_matches_prefix stringappend_15470 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15480,stringappend_15490) => + (stringappend_15480, stringappend_15490) )) in - (let stringappend_14960 = (string_drop stringappend_14930 - stringappend_14950) in + (let stringappend_15500 = (string_drop stringappend_15470 + stringappend_15490) in (case - (case ((sep_matches_prefix stringappend_14960)) of - Some (stringappend_14970,stringappend_14980) => - (stringappend_14970, stringappend_14980) + (case ((sep_matches_prefix stringappend_15500)) of + Some (stringappend_15510,stringappend_15520) => + (stringappend_15510, stringappend_15520) ) of - (_, stringappend_14980) => - (let stringappend_14990 = (string_drop stringappend_14960 - stringappend_14980) in - (let (rs1, stringappend_15010) = - ((case ((reg_name_matches_prefix stringappend_14990 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15000,stringappend_15010) => - (stringappend_15000, stringappend_15010) + (_, stringappend_15520) => + (let stringappend_15530 = (string_drop stringappend_15500 + stringappend_15520) in + (let (rs1, stringappend_15550) = + ((case ((reg_name_matches_prefix stringappend_15530 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15540,stringappend_15550) => + (stringappend_15540, stringappend_15550) )) in - if(((string_drop stringappend_14990 stringappend_15010)) = + if(((string_drop stringappend_15530 stringappend_15550)) = ('''')) then (LOADRES (aq,rl,rs1,size1,rd)) else undefined)) )))) )))))))) else - if (((((string_startswith stringappend_10760 (''sc.''))) + if (((((string_startswith stringappend_11120 (''sc.''))) \<and> - ((let stringappend_15030 = - (string_drop stringappend_10760 + ((let stringappend_15570 = + (string_drop stringappend_11120 ((string_length (''sc.'')))) in if ((case ((maybe_aq_matches_prefix - stringappend_15030)) of - Some (stringappend_15040,stringappend_15050) => - (let stringappend_15060 = - (string_drop stringappend_15030 - stringappend_15050) in + stringappend_15570)) of + Some (stringappend_15580,stringappend_15590) => + (let stringappend_15600 = + (string_drop stringappend_15570 + stringappend_15590) in if ((case ((maybe_rl_matches_prefix - stringappend_15060)) of - Some (stringappend_15070,stringappend_15080) => - (let stringappend_15090 = - (string_drop stringappend_15060 - stringappend_15080) in + stringappend_15600)) of + Some (stringappend_15610,stringappend_15620) => + (let stringappend_15630 = + (string_drop stringappend_15600 + stringappend_15620) in if ((case ((size_mnemonic_matches_prefix - stringappend_15090)) of - Some (stringappend_15100,stringappend_15110) => - (let stringappend_15120 = - (string_drop stringappend_15090 - stringappend_15110) in + stringappend_15630)) of + Some (stringappend_15640,stringappend_15650) => + (let stringappend_15660 = + (string_drop stringappend_15630 + stringappend_15650) in if ((case ((spc_matches_prefix - stringappend_15120)) of - Some (stringappend_15130,stringappend_15140) => - (let stringappend_15150 = + stringappend_15660)) of + Some (stringappend_15670,stringappend_15680) => + (let stringappend_15690 = (string_drop - stringappend_15120 - stringappend_15140) in + stringappend_15660 + stringappend_15680) in if ((case ((reg_name_matches_prefix - stringappend_15150 + stringappend_15690 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15160,stringappend_15170) => - (let stringappend_15180 = + Some (stringappend_15700,stringappend_15710) => + (let stringappend_15720 = (string_drop - stringappend_15150 - stringappend_15170) in + stringappend_15690 + stringappend_15710) in if ((case ((sep_matches_prefix - stringappend_15180)) of - Some (stringappend_15190,stringappend_15200) => - (let stringappend_15210 = + stringappend_15720)) of + Some (stringappend_15730,stringappend_15740) => + (let stringappend_15750 = (string_drop - stringappend_15180 - stringappend_15200) in + stringappend_15720 + stringappend_15740) in if ((case ((reg_name_matches_prefix - stringappend_15210 + stringappend_15750 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15220,stringappend_15230) => - (let stringappend_15240 = + Some (stringappend_15760,stringappend_15770) => + (let stringappend_15780 = (string_drop - stringappend_15210 - stringappend_15230) in + stringappend_15750 + stringappend_15770) in if ((case ((sep_matches_prefix - stringappend_15240)) of + stringappend_15780)) of Some - (stringappend_15250,stringappend_15260) => - (let stringappend_15270 = + (stringappend_15790,stringappend_15800) => + (let stringappend_15810 = (string_drop - stringappend_15240 - stringappend_15260) in + stringappend_15780 + stringappend_15800) in if ((case ((reg_name_matches_prefix - stringappend_15270 + stringappend_15810 :: (( 5 Word.word * ii)) option)) of Some - (stringappend_15280,stringappend_15290) => + (stringappend_15820,stringappend_15830) => if(((string_drop - stringappend_15270 - stringappend_15290)) + stringappend_15810 + stringappend_15830)) = ('''')) then True else @@ -16400,166 +16567,166 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False) | None => False )) then True else False))))) then - (let stringappend_15030 = (string_drop stringappend_10760 + (let stringappend_15570 = (string_drop stringappend_11120 ((string_length (''sc.'')))) in - (let (aq, stringappend_15050) = - ((case ((maybe_aq_matches_prefix stringappend_15030)) of - Some (stringappend_15040,stringappend_15050) => - (stringappend_15040, stringappend_15050) + (let (aq, stringappend_15590) = + ((case ((maybe_aq_matches_prefix stringappend_15570)) of + Some (stringappend_15580,stringappend_15590) => + (stringappend_15580, stringappend_15590) )) in - (let stringappend_15060 = (string_drop stringappend_15030 - stringappend_15050) in - (let (rl, stringappend_15080) = - ((case ((maybe_rl_matches_prefix stringappend_15060)) of - Some (stringappend_15070,stringappend_15080) => - (stringappend_15070, stringappend_15080) + (let stringappend_15600 = (string_drop stringappend_15570 + stringappend_15590) in + (let (rl, stringappend_15620) = + ((case ((maybe_rl_matches_prefix stringappend_15600)) of + Some (stringappend_15610,stringappend_15620) => + (stringappend_15610, stringappend_15620) )) in - (let stringappend_15090 = (string_drop stringappend_15060 - stringappend_15080) in - (let (size1, stringappend_15110) = + (let stringappend_15630 = (string_drop stringappend_15600 + stringappend_15620) in + (let (size1, stringappend_15650) = ((case ((size_mnemonic_matches_prefix - stringappend_15090)) of - Some (stringappend_15100,stringappend_15110) => - (stringappend_15100, stringappend_15110) + stringappend_15630)) of + Some (stringappend_15640,stringappend_15650) => + (stringappend_15640, stringappend_15650) )) in - (let stringappend_15120 = (string_drop stringappend_15090 - stringappend_15110) in + (let stringappend_15660 = (string_drop stringappend_15630 + stringappend_15650) in (case - (case ((spc_matches_prefix stringappend_15120)) of - Some (stringappend_15130,stringappend_15140) => - (stringappend_15130, stringappend_15140) + (case ((spc_matches_prefix stringappend_15660)) of + Some (stringappend_15670,stringappend_15680) => + (stringappend_15670, stringappend_15680) ) of - (_, stringappend_15140) => - (let stringappend_15150 = (string_drop stringappend_15120 - stringappend_15140) in - (let (rd, stringappend_15170) = - ((case ((reg_name_matches_prefix stringappend_15150 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15160,stringappend_15170) => - (stringappend_15160, stringappend_15170) + (_, stringappend_15680) => + (let stringappend_15690 = (string_drop stringappend_15660 + stringappend_15680) in + (let (rd, stringappend_15710) = + ((case ((reg_name_matches_prefix stringappend_15690 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15700,stringappend_15710) => + (stringappend_15700, stringappend_15710) )) in - (let stringappend_15180 = (string_drop stringappend_15150 - stringappend_15170) in + (let stringappend_15720 = (string_drop stringappend_15690 + stringappend_15710) in (case - (case ((sep_matches_prefix stringappend_15180)) of - Some (stringappend_15190,stringappend_15200) => - (stringappend_15190, stringappend_15200) + (case ((sep_matches_prefix stringappend_15720)) of + Some (stringappend_15730,stringappend_15740) => + (stringappend_15730, stringappend_15740) ) of - (_, stringappend_15200) => - (let stringappend_15210 = (string_drop stringappend_15180 - stringappend_15200) in - (let (rs1, stringappend_15230) = - ((case ((reg_name_matches_prefix stringappend_15210 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15220,stringappend_15230) => - (stringappend_15220, stringappend_15230) + (_, stringappend_15740) => + (let stringappend_15750 = (string_drop stringappend_15720 + stringappend_15740) in + (let (rs1, stringappend_15770) = + ((case ((reg_name_matches_prefix stringappend_15750 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15760,stringappend_15770) => + (stringappend_15760, stringappend_15770) )) in - (let stringappend_15240 = (string_drop stringappend_15210 - stringappend_15230) in + (let stringappend_15780 = (string_drop stringappend_15750 + stringappend_15770) in (case - (case ((sep_matches_prefix stringappend_15240)) of - Some (stringappend_15250,stringappend_15260) => - (stringappend_15250, stringappend_15260) + (case ((sep_matches_prefix stringappend_15780)) of + Some (stringappend_15790,stringappend_15800) => + (stringappend_15790, stringappend_15800) ) of - (_, stringappend_15260) => - (let stringappend_15270 = (string_drop stringappend_15240 - stringappend_15260) in - (let (rs2, stringappend_15290) = - ((case ((reg_name_matches_prefix stringappend_15270 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15280,stringappend_15290) => - (stringappend_15280, stringappend_15290) + (_, stringappend_15800) => + (let stringappend_15810 = (string_drop stringappend_15780 + stringappend_15800) in + (let (rs2, stringappend_15830) = + ((case ((reg_name_matches_prefix stringappend_15810 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_15820,stringappend_15830) => + (stringappend_15820, stringappend_15830) )) in - if(((string_drop stringappend_15270 stringappend_15290)) = + if(((string_drop stringappend_15810 stringappend_15830)) = ('''')) then (STORECON (aq,rl,rs2,rs1,size1,rd)) else undefined)) )))) )))) )))))))) else if ((case ((amo_mnemonic_matches_prefix - stringappend_10760)) of - Some (stringappend_15310,stringappend_15320) => - (let stringappend_15330 = (string_drop - stringappend_10760 - stringappend_15320) in - if (((((string_startswith stringappend_15330 (''.''))) + stringappend_11120)) of + Some (stringappend_15850,stringappend_15860) => + (let stringappend_15870 = (string_drop + stringappend_11120 + stringappend_15860) in + if (((((string_startswith stringappend_15870 (''.''))) \<and> - ((let stringappend_15340 = (string_drop - stringappend_15330 + ((let stringappend_15880 = (string_drop + stringappend_15870 ((string_length (''.'')))) in if ((case ((size_mnemonic_matches_prefix - stringappend_15340)) of - Some (stringappend_15350,stringappend_15360) => - (let stringappend_15370 = - (string_drop stringappend_15340 - stringappend_15360) in + stringappend_15880)) of + Some (stringappend_15890,stringappend_15900) => + (let stringappend_15910 = + (string_drop stringappend_15880 + stringappend_15900) in if ((case ((maybe_aq_matches_prefix - stringappend_15370)) of - Some (stringappend_15380,stringappend_15390) => - (let stringappend_15400 = - (string_drop stringappend_15370 - stringappend_15390) in + stringappend_15910)) of + Some (stringappend_15920,stringappend_15930) => + (let stringappend_15940 = + (string_drop stringappend_15910 + stringappend_15930) in if ((case ((maybe_rl_matches_prefix - stringappend_15400)) of - Some (stringappend_15410,stringappend_15420) => - (let stringappend_15430 = + stringappend_15940)) of + Some (stringappend_15950,stringappend_15960) => + (let stringappend_15970 = (string_drop - stringappend_15400 - stringappend_15420) in + stringappend_15940 + stringappend_15960) in if ((case ((spc_matches_prefix - stringappend_15430)) of - Some (stringappend_15440,stringappend_15450) => - (let stringappend_15460 = + stringappend_15970)) of + Some (stringappend_15980,stringappend_15990) => + (let stringappend_16000 = (string_drop - stringappend_15430 - stringappend_15450) in + stringappend_15970 + stringappend_15990) in if ((case ((reg_name_matches_prefix - stringappend_15460 + stringappend_16000 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15470,stringappend_15480) => - (let stringappend_15490 = + Some (stringappend_16010,stringappend_16020) => + (let stringappend_16030 = (string_drop - stringappend_15460 - stringappend_15480) in + stringappend_16000 + stringappend_16020) in if ((case ((sep_matches_prefix - stringappend_15490)) of - Some (stringappend_15500,stringappend_15510) => - (let stringappend_15520 = + stringappend_16030)) of + Some (stringappend_16040,stringappend_16050) => + (let stringappend_16060 = (string_drop - stringappend_15490 - stringappend_15510) in + stringappend_16030 + stringappend_16050) in if ((case (( reg_name_matches_prefix - stringappend_15520 + stringappend_16060 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15530,stringappend_15540) => - (let stringappend_15550 = + Some (stringappend_16070,stringappend_16080) => + (let stringappend_16090 = (string_drop - stringappend_15520 - stringappend_15540) in + stringappend_16060 + stringappend_16080) in if ((case ((sep_matches_prefix - stringappend_15550)) of + stringappend_16090)) of Some - (stringappend_15560,stringappend_15570) => + (stringappend_16100,stringappend_16110) => (let - stringappend_15580 = + stringappend_16120 = (string_drop - stringappend_15550 - stringappend_15570) in + stringappend_16090 + stringappend_16110) in if ((case ( ( reg_name_matches_prefix - stringappend_15580 + stringappend_16120 :: (( 5 Word.word * ii)) option)) of Some - (stringappend_15590,stringappend_15600) => + (stringappend_16130,stringappend_16140) => if ( ( ( string_drop - stringappend_15580 - stringappend_15600)) + stringappend_16120 + stringappend_16140)) = ('''')) then True else @@ -16595,158 +16762,158 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where True else False) | None => False )) then - (let (op1, stringappend_15320) = + (let (op1, stringappend_15860) = ((case ((amo_mnemonic_matches_prefix - stringappend_10760)) of - Some (stringappend_15310,stringappend_15320) => - (stringappend_15310, stringappend_15320) + stringappend_11120)) of + Some (stringappend_15850,stringappend_15860) => + (stringappend_15850, stringappend_15860) )) in - (let stringappend_15330 = (string_drop - stringappend_10760 - stringappend_15320) in - (let stringappend_15340 = (string_drop - stringappend_15330 + (let stringappend_15870 = (string_drop + stringappend_11120 + stringappend_15860) in + (let stringappend_15880 = (string_drop + stringappend_15870 ((string_length (''.'')))) in - (let (width, stringappend_15360) = + (let (width, stringappend_15900) = ((case ((size_mnemonic_matches_prefix - stringappend_15340)) of - Some (stringappend_15350,stringappend_15360) => - (stringappend_15350, stringappend_15360) + stringappend_15880)) of + Some (stringappend_15890,stringappend_15900) => + (stringappend_15890, stringappend_15900) )) in - (let stringappend_15370 = (string_drop - stringappend_15340 - stringappend_15360) in - (let (aq, stringappend_15390) = + (let stringappend_15910 = (string_drop + stringappend_15880 + stringappend_15900) in + (let (aq, stringappend_15930) = ((case ((maybe_aq_matches_prefix - stringappend_15370)) of - Some (stringappend_15380,stringappend_15390) => - (stringappend_15380, stringappend_15390) + stringappend_15910)) of + Some (stringappend_15920,stringappend_15930) => + (stringappend_15920, stringappend_15930) )) in - (let stringappend_15400 = (string_drop - stringappend_15370 - stringappend_15390) in - (let (rl, stringappend_15420) = + (let stringappend_15940 = (string_drop + stringappend_15910 + stringappend_15930) in + (let (rl, stringappend_15960) = ((case ((maybe_rl_matches_prefix - stringappend_15400)) of - Some (stringappend_15410,stringappend_15420) => - (stringappend_15410, stringappend_15420) + stringappend_15940)) of + Some (stringappend_15950,stringappend_15960) => + (stringappend_15950, stringappend_15960) )) in - (let stringappend_15430 = (string_drop - stringappend_15400 - stringappend_15420) in + (let stringappend_15970 = (string_drop + stringappend_15940 + stringappend_15960) in (case - (case ((spc_matches_prefix stringappend_15430)) of - Some (stringappend_15440,stringappend_15450) => - (stringappend_15440, stringappend_15450) + (case ((spc_matches_prefix stringappend_15970)) of + Some (stringappend_15980,stringappend_15990) => + (stringappend_15980, stringappend_15990) ) of - (_, stringappend_15450) => - (let stringappend_15460 = (string_drop - stringappend_15430 - stringappend_15450) in - (let (rd, stringappend_15480) = + (_, stringappend_15990) => + (let stringappend_16000 = (string_drop + stringappend_15970 + stringappend_15990) in + (let (rd, stringappend_16020) = ((case ((reg_name_matches_prefix - stringappend_15460 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15470,stringappend_15480) => - (stringappend_15470, stringappend_15480) + stringappend_16000 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_16010,stringappend_16020) => + (stringappend_16010, stringappend_16020) )) in - (let stringappend_15490 = (string_drop - stringappend_15460 - stringappend_15480) in + (let stringappend_16030 = (string_drop + stringappend_16000 + stringappend_16020) in (case - (case ((sep_matches_prefix stringappend_15490)) of - Some (stringappend_15500,stringappend_15510) => - (stringappend_15500, stringappend_15510) + (case ((sep_matches_prefix stringappend_16030)) of + Some (stringappend_16040,stringappend_16050) => + (stringappend_16040, stringappend_16050) ) of - (_, stringappend_15510) => - (let stringappend_15520 = (string_drop - stringappend_15490 - stringappend_15510) in - (let (rs1, stringappend_15540) = + (_, stringappend_16050) => + (let stringappend_16060 = (string_drop + stringappend_16030 + stringappend_16050) in + (let (rs1, stringappend_16080) = ((case ((reg_name_matches_prefix - stringappend_15520 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15530,stringappend_15540) => - (stringappend_15530, stringappend_15540) + stringappend_16060 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_16070,stringappend_16080) => + (stringappend_16070, stringappend_16080) )) in - (let stringappend_15550 = (string_drop - stringappend_15520 - stringappend_15540) in + (let stringappend_16090 = (string_drop + stringappend_16060 + stringappend_16080) in (case - (case ((sep_matches_prefix stringappend_15550)) of - Some (stringappend_15560,stringappend_15570) => - (stringappend_15560, stringappend_15570) + (case ((sep_matches_prefix stringappend_16090)) of + Some (stringappend_16100,stringappend_16110) => + (stringappend_16100, stringappend_16110) ) of - (_, stringappend_15570) => - (let stringappend_15580 = (string_drop - stringappend_15550 - stringappend_15570) in - (let (rs2, stringappend_15600) = + (_, stringappend_16110) => + (let stringappend_16120 = (string_drop + stringappend_16090 + stringappend_16110) in + (let (rs2, stringappend_16140) = ((case ((reg_name_matches_prefix - stringappend_15580 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15590,stringappend_15600) => - (stringappend_15590, stringappend_15600) + stringappend_16120 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_16130,stringappend_16140) => + (stringappend_16130, stringappend_16140) )) in - if(((string_drop stringappend_15580 stringappend_15600)) + if(((string_drop stringappend_16120 stringappend_16140)) = ('''')) then (AMO (op1,aq,rl,rs2,rs1,width,rd)) else undefined)) )))) )))) )))))))))) else if ((case ((csr_mnemonic_matches_prefix - stringappend_10760)) of - Some (stringappend_15620,stringappend_15630) => - (let stringappend_15640 = (string_drop - stringappend_10760 - stringappend_15630) in - if (((((string_startswith stringappend_15640 (''i''))) + stringappend_11120)) of + Some (stringappend_16160,stringappend_16170) => + (let stringappend_16180 = (string_drop + stringappend_11120 + stringappend_16170) in + if (((((string_startswith stringappend_16180 (''i''))) \<and> - ((let stringappend_15650 = (string_drop - stringappend_15640 + ((let stringappend_16190 = (string_drop + stringappend_16180 ((string_length (''i'')))) in if ((case ((spc_matches_prefix - stringappend_15650)) of - Some (stringappend_15660,stringappend_15670) => - (let stringappend_15680 = - (string_drop stringappend_15650 - stringappend_15670) in + stringappend_16190)) of + Some (stringappend_16200,stringappend_16210) => + (let stringappend_16220 = + (string_drop stringappend_16190 + stringappend_16210) in if ((case ((reg_name_matches_prefix - stringappend_15680 + stringappend_16220 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15690,stringappend_15700) => - (let stringappend_15710 = + Some (stringappend_16230,stringappend_16240) => + (let stringappend_16250 = (string_drop - stringappend_15680 - stringappend_15700) in + stringappend_16220 + stringappend_16240) in if ((case ((sep_matches_prefix - stringappend_15710)) of - Some (stringappend_15720,stringappend_15730) => - (let stringappend_15740 = + stringappend_16250)) of + Some (stringappend_16260,stringappend_16270) => + (let stringappend_16280 = (string_drop - stringappend_15710 - stringappend_15730) in + stringappend_16250 + stringappend_16270) in if ((case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_15740 + stringappend_16280 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15750,stringappend_15760) => - (let stringappend_15770 = + Some (stringappend_16290,stringappend_16300) => + (let stringappend_16310 = (string_drop - stringappend_15740 - stringappend_15760) in + stringappend_16280 + stringappend_16300) in if ((case ((sep_matches_prefix - stringappend_15770)) of - Some (stringappend_15780,stringappend_15790) => - (let stringappend_15800 = + stringappend_16310)) of + Some (stringappend_16320,stringappend_16330) => + (let stringappend_16340 = (string_drop - stringappend_15770 - stringappend_15790) in + stringappend_16310 + stringappend_16330) in if ((case ((csr_name_map_matches_prefix - stringappend_15800 + stringappend_16340 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_15810,stringappend_15820) => + Some (stringappend_16350,stringappend_16360) => if(((string_drop - stringappend_15800 - stringappend_15820)) + stringappend_16340 + stringappend_16360)) = ('''')) then True else False @@ -16768,122 +16935,122 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where True else False) | None => False )) then - (let (op1, stringappend_15630) = + (let (op1, stringappend_16170) = ((case ((csr_mnemonic_matches_prefix - stringappend_10760)) of - Some (stringappend_15620,stringappend_15630) => - (stringappend_15620, stringappend_15630) + stringappend_11120)) of + Some (stringappend_16160,stringappend_16170) => + (stringappend_16160, stringappend_16170) )) in - (let stringappend_15640 = (string_drop - stringappend_10760 - stringappend_15630) in - (let stringappend_15650 = (string_drop - stringappend_15640 + (let stringappend_16180 = (string_drop + stringappend_11120 + stringappend_16170) in + (let stringappend_16190 = (string_drop + stringappend_16180 ((string_length (''i'')))) in (case - (case ((spc_matches_prefix stringappend_15650)) of - Some (stringappend_15660,stringappend_15670) => - (stringappend_15660, stringappend_15670) + (case ((spc_matches_prefix stringappend_16190)) of + Some (stringappend_16200,stringappend_16210) => + (stringappend_16200, stringappend_16210) ) of - (_, stringappend_15670) => - (let stringappend_15680 = (string_drop - stringappend_15650 - stringappend_15670) in - (let (rd, stringappend_15700) = + (_, stringappend_16210) => + (let stringappend_16220 = (string_drop + stringappend_16190 + stringappend_16210) in + (let (rd, stringappend_16240) = ((case ((reg_name_matches_prefix - stringappend_15680 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15690,stringappend_15700) => - (stringappend_15690, stringappend_15700) + stringappend_16220 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_16230,stringappend_16240) => + (stringappend_16230, stringappend_16240) )) in - (let stringappend_15710 = (string_drop - stringappend_15680 - stringappend_15700) in + (let stringappend_16250 = (string_drop + stringappend_16220 + stringappend_16240) in (case - (case ((sep_matches_prefix stringappend_15710)) of - Some (stringappend_15720,stringappend_15730) => - (stringappend_15720, stringappend_15730) + (case ((sep_matches_prefix stringappend_16250)) of + Some (stringappend_16260,stringappend_16270) => + (stringappend_16260, stringappend_16270) ) of - (_, stringappend_15730) => - (let stringappend_15740 = (string_drop - stringappend_15710 - stringappend_15730) in - (let (rs1, stringappend_15760) = + (_, stringappend_16270) => + (let stringappend_16280 = (string_drop + stringappend_16250 + stringappend_16270) in + (let (rs1, stringappend_16300) = ((case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_15740 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15750,stringappend_15760) => - (stringappend_15750, stringappend_15760) + stringappend_16280 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_16290,stringappend_16300) => + (stringappend_16290, stringappend_16300) )) in - (let stringappend_15770 = (string_drop - stringappend_15740 - stringappend_15760) in + (let stringappend_16310 = (string_drop + stringappend_16280 + stringappend_16300) in (case - (case ((sep_matches_prefix stringappend_15770)) of - Some (stringappend_15780,stringappend_15790) => - (stringappend_15780, stringappend_15790) + (case ((sep_matches_prefix stringappend_16310)) of + Some (stringappend_16320,stringappend_16330) => + (stringappend_16320, stringappend_16330) ) of - (_, stringappend_15790) => - (let stringappend_15800 = (string_drop - stringappend_15770 - stringappend_15790) in - (let (csr, stringappend_15820) = + (_, stringappend_16330) => + (let stringappend_16340 = (string_drop + stringappend_16310 + stringappend_16330) in + (let (csr, stringappend_16360) = ((case ((csr_name_map_matches_prefix - stringappend_15800 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_15810,stringappend_15820) => - (stringappend_15810, stringappend_15820) + stringappend_16340 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_16350,stringappend_16360) => + (stringappend_16350, stringappend_16360) )) in - if(((string_drop stringappend_15800 stringappend_15820)) + if(((string_drop stringappend_16340 stringappend_16360)) = ('''')) then (CSR (csr,rs1,rd,True,op1)) else undefined)) )))) )))) )))) else if ((case ((csr_mnemonic_matches_prefix - stringappend_10760)) of - Some (stringappend_15840,stringappend_15850) => - (let stringappend_15860 = (string_drop - stringappend_10760 - stringappend_15850) in - if ((case ((spc_matches_prefix stringappend_15860)) of - Some (stringappend_15870,stringappend_15880) => - (let stringappend_15890 = (string_drop - stringappend_15860 - stringappend_15880) in + stringappend_11120)) of + Some (stringappend_16380,stringappend_16390) => + (let stringappend_16400 = (string_drop + stringappend_11120 + stringappend_16390) in + if ((case ((spc_matches_prefix stringappend_16400)) of + Some (stringappend_16410,stringappend_16420) => + (let stringappend_16430 = (string_drop + stringappend_16400 + stringappend_16420) in if ((case ((reg_name_matches_prefix - stringappend_15890 + stringappend_16430 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15900,stringappend_15910) => - (let stringappend_15920 = (string_drop - stringappend_15890 - stringappend_15910) in + Some (stringappend_16440,stringappend_16450) => + (let stringappend_16460 = (string_drop + stringappend_16430 + stringappend_16450) in if ((case ((sep_matches_prefix - stringappend_15920)) of - Some (stringappend_15930,stringappend_15940) => - (let stringappend_15950 = - (string_drop stringappend_15920 - stringappend_15940) in + stringappend_16460)) of + Some (stringappend_16470,stringappend_16480) => + (let stringappend_16490 = + (string_drop stringappend_16460 + stringappend_16480) in if ((case ((reg_name_matches_prefix - stringappend_15950 + stringappend_16490 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15960,stringappend_15970) => - (let stringappend_15980 = + Some (stringappend_16500,stringappend_16510) => + (let stringappend_16520 = (string_drop - stringappend_15950 - stringappend_15970) in + stringappend_16490 + stringappend_16510) in if ((case ((sep_matches_prefix - stringappend_15980)) of - Some (stringappend_15990,stringappend_16000) => - (let stringappend_16010 = + stringappend_16520)) of + Some (stringappend_16530,stringappend_16540) => + (let stringappend_16550 = (string_drop - stringappend_15980 - stringappend_16000) in + stringappend_16520 + stringappend_16540) in if ((case ((csr_name_map_matches_prefix - stringappend_16010 + stringappend_16550 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_16020,stringappend_16030) => + Some (stringappend_16560,stringappend_16570) => if(((string_drop - stringappend_16010 - stringappend_16030)) + stringappend_16550 + stringappend_16570)) = ('''')) then True else False | None => False @@ -16901,94 +17068,94 @@ definition assembly_backwards :: " string \<Rightarrow> ast " where )) then True else False) | None => False )) then - (let (op1, stringappend_15850) = + (let (op1, stringappend_16390) = ((case ((csr_mnemonic_matches_prefix - stringappend_10760)) of - Some (stringappend_15840,stringappend_15850) => - (stringappend_15840, stringappend_15850) + stringappend_11120)) of + Some (stringappend_16380,stringappend_16390) => + (stringappend_16380, stringappend_16390) )) in - (let stringappend_15860 = (string_drop - stringappend_10760 - stringappend_15850) in + (let stringappend_16400 = (string_drop + stringappend_11120 + stringappend_16390) in (case - (case ((spc_matches_prefix stringappend_15860)) of - Some (stringappend_15870,stringappend_15880) => - (stringappend_15870, stringappend_15880) + (case ((spc_matches_prefix stringappend_16400)) of + Some (stringappend_16410,stringappend_16420) => + (stringappend_16410, stringappend_16420) ) of - (_, stringappend_15880) => - (let stringappend_15890 = (string_drop - stringappend_15860 - stringappend_15880) in - (let (rd, stringappend_15910) = + (_, stringappend_16420) => + (let stringappend_16430 = (string_drop + stringappend_16400 + stringappend_16420) in + (let (rd, stringappend_16450) = ((case ((reg_name_matches_prefix - stringappend_15890 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15900,stringappend_15910) => - (stringappend_15900, stringappend_15910) + stringappend_16430 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_16440,stringappend_16450) => + (stringappend_16440, stringappend_16450) )) in - (let stringappend_15920 = (string_drop - stringappend_15890 - stringappend_15910) in + (let stringappend_16460 = (string_drop + stringappend_16430 + stringappend_16450) in (case - (case ((sep_matches_prefix stringappend_15920)) of - Some (stringappend_15930,stringappend_15940) => - (stringappend_15930, stringappend_15940) + (case ((sep_matches_prefix stringappend_16460)) of + Some (stringappend_16470,stringappend_16480) => + (stringappend_16470, stringappend_16480) ) of - (_, stringappend_15940) => - (let stringappend_15950 = (string_drop - stringappend_15920 - stringappend_15940) in - (let (rs1, stringappend_15970) = + (_, stringappend_16480) => + (let stringappend_16490 = (string_drop + stringappend_16460 + stringappend_16480) in + (let (rs1, stringappend_16510) = ((case ((reg_name_matches_prefix - stringappend_15950 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_15960,stringappend_15970) => - (stringappend_15960, stringappend_15970) + stringappend_16490 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_16500,stringappend_16510) => + (stringappend_16500, stringappend_16510) )) in - (let stringappend_15980 = (string_drop - stringappend_15950 - stringappend_15970) in + (let stringappend_16520 = (string_drop + stringappend_16490 + stringappend_16510) in (case - (case ((sep_matches_prefix stringappend_15980)) of - Some (stringappend_15990,stringappend_16000) => - (stringappend_15990, stringappend_16000) + (case ((sep_matches_prefix stringappend_16520)) of + Some (stringappend_16530,stringappend_16540) => + (stringappend_16530, stringappend_16540) ) of - (_, stringappend_16000) => - (let stringappend_16010 = (string_drop - stringappend_15980 - stringappend_16000) in - (let (csr, stringappend_16030) = + (_, stringappend_16540) => + (let stringappend_16550 = (string_drop + stringappend_16520 + stringappend_16540) in + (let (csr, stringappend_16570) = ((case ((csr_name_map_matches_prefix - stringappend_16010 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_16020,stringappend_16030) => - (stringappend_16020, stringappend_16030) + stringappend_16550 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_16560,stringappend_16570) => + (stringappend_16560, stringappend_16570) )) in - if(((string_drop stringappend_16010 - stringappend_16030)) = ('''')) then + if(((string_drop stringappend_16550 + stringappend_16570)) = ('''')) then (CSR (csr,rs1,rd,False,op1)) else undefined)) )))) )))) ))) else - (let stringappend_16050 = (string_drop - stringappend_10760 + (let stringappend_16590 = (string_drop + stringappend_11120 ((string_length (''illegal'')))) in (case - (case ((spc_matches_prefix stringappend_16050)) of - Some (stringappend_16060,stringappend_16070) => - (stringappend_16060, stringappend_16070) + (case ((spc_matches_prefix stringappend_16590)) of + Some (stringappend_16600,stringappend_16610) => + (stringappend_16600, stringappend_16610) ) of - (_, stringappend_16070) => - (let stringappend_16080 = (string_drop - stringappend_16050 - stringappend_16070) in - (let (s, stringappend_16100) = + (_, stringappend_16610) => + (let stringappend_16620 = (string_drop + stringappend_16590 + stringappend_16610) in + (let (s, stringappend_16640) = ((case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_16080 :: (( 32 Word.word * ii)) option)) of - Some (stringappend_16090,stringappend_16100) => - (stringappend_16090, stringappend_16100) + stringappend_16620 :: (( 32 Word.word * ii)) option)) of + Some (stringappend_16630,stringappend_16640) => + (stringappend_16630, stringappend_16640) )) in - if(((string_drop stringappend_16080 - stringappend_16100)) = ('''')) then + if(((string_drop stringappend_16620 + stringappend_16640)) = ('''')) then (ILLEGAL s) else undefined)) ))))))))))" @@ -17008,6 +17175,7 @@ fun assembly_forwards_matches :: " ast \<Rightarrow> bool " where |" assembly_forwards_matches (ADDIW (imm,rs1,rd)) = ( True )" |" assembly_forwards_matches (SHIFTW (shamt,rs1,rd,op1)) = ( True )" |" assembly_forwards_matches (RTYPEW (rs2,rs1,rd,op1)) = ( True )" +|" assembly_forwards_matches (SHIFTIWOP (shamt,rs1,rd,op1)) = ( True )" |" assembly_forwards_matches (MUL (rs2,rs1,rd,high,signed1,signed2)) = ( True )" |" assembly_forwards_matches (DIV (rs2,rs1,rd,s)) = ( True )" |" assembly_forwards_matches (REM (rs2,rs1,rd,s)) = ( True )" @@ -17035,24 +17203,24 @@ fun assembly_forwards_matches :: " ast \<Rightarrow> bool " where definition assembly_backwards_matches :: " string \<Rightarrow> bool " where " assembly_backwards_matches arg0 = ( - (let stringappend_5410 = arg0 in - if ((case ((utype_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_5420,stringappend_5430) => - (let stringappend_5440 = (string_drop stringappend_5410 stringappend_5430) in - if ((case ((spc_matches_prefix stringappend_5440)) of - Some (stringappend_5450,stringappend_5460) => - (let stringappend_5470 = (string_drop stringappend_5440 stringappend_5460) in - if ((case ((reg_name_matches_prefix stringappend_5470 :: (( 5 Word.word * ii))option)) of - Some (stringappend_5480,stringappend_5490) => - (let stringappend_5500 = (string_drop stringappend_5470 stringappend_5490) in - if ((case ((sep_matches_prefix stringappend_5500)) of - Some (stringappend_5510,stringappend_5520) => - (let stringappend_5530 = (string_drop stringappend_5500 stringappend_5520) in + (let stringappend_5590 = arg0 in + if ((case ((utype_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_5600,stringappend_5610) => + (let stringappend_5620 = (string_drop stringappend_5590 stringappend_5610) in + if ((case ((spc_matches_prefix stringappend_5620)) of + Some (stringappend_5630,stringappend_5640) => + (let stringappend_5650 = (string_drop stringappend_5620 stringappend_5640) in + if ((case ((reg_name_matches_prefix stringappend_5650 :: (( 5 Word.word * ii))option)) of + Some (stringappend_5660,stringappend_5670) => + (let stringappend_5680 = (string_drop stringappend_5650 stringappend_5670) in + if ((case ((sep_matches_prefix stringappend_5680)) of + Some (stringappend_5690,stringappend_5700) => + (let stringappend_5710 = (string_drop stringappend_5680 stringappend_5700) in if ((case ((hex_bits_20_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5530 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5710 :: (( 20 Word.word * ii))option)) of - Some (stringappend_5540,stringappend_5550) => - if(((string_drop stringappend_5530 stringappend_5550)) = ('''')) then + Some (stringappend_5720,stringappend_5730) => + if(((string_drop stringappend_5710 stringappend_5730)) = ('''')) then True else False | None => False )) then @@ -17072,55 +17240,55 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let (op1, stringappend_5430) = - ((case ((utype_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_5420,stringappend_5430) => (stringappend_5420, stringappend_5430) + (let (op1, stringappend_5610) = + ((case ((utype_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_5600,stringappend_5610) => (stringappend_5600, stringappend_5610) )) in - (let stringappend_5440 = (string_drop stringappend_5410 stringappend_5430) in + (let stringappend_5620 = (string_drop stringappend_5590 stringappend_5610) in (case - (case ((spc_matches_prefix stringappend_5440)) of - Some (stringappend_5450,stringappend_5460) => (stringappend_5450, stringappend_5460) + (case ((spc_matches_prefix stringappend_5620)) of + Some (stringappend_5630,stringappend_5640) => (stringappend_5630, stringappend_5640) ) of - (_, stringappend_5460) => - (let stringappend_5470 = (string_drop stringappend_5440 stringappend_5460) in - (let (rd, stringappend_5490) = - ((case ((reg_name_matches_prefix stringappend_5470 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_5480,stringappend_5490) => (stringappend_5480, stringappend_5490) + (_, stringappend_5640) => + (let stringappend_5650 = (string_drop stringappend_5620 stringappend_5640) in + (let (rd, stringappend_5670) = + ((case ((reg_name_matches_prefix stringappend_5650 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5660,stringappend_5670) => (stringappend_5660, stringappend_5670) )) in - (let stringappend_5500 = (string_drop stringappend_5470 stringappend_5490) in + (let stringappend_5680 = (string_drop stringappend_5650 stringappend_5670) in (case - (case ((sep_matches_prefix stringappend_5500)) of - Some (stringappend_5510,stringappend_5520) => (stringappend_5510, stringappend_5520) + (case ((sep_matches_prefix stringappend_5680)) of + Some (stringappend_5690,stringappend_5700) => (stringappend_5690, stringappend_5700) ) of - (_, stringappend_5520) => - (let stringappend_5530 = (string_drop stringappend_5500 stringappend_5520) in - (let (imm, stringappend_5550) = + (_, stringappend_5700) => + (let stringappend_5710 = (string_drop stringappend_5680 stringappend_5700) in + (let (imm, stringappend_5730) = ((case ((hex_bits_20_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_5530 :: (( 20 Word.word * ii)) option)) of - Some (stringappend_5540,stringappend_5550) => (stringappend_5540, stringappend_5550) + stringappend_5710 :: (( 20 Word.word * ii)) option)) of + Some (stringappend_5720,stringappend_5730) => (stringappend_5720, stringappend_5730) )) in - if(((string_drop stringappend_5530 stringappend_5550)) = ('''')) then + if(((string_drop stringappend_5710 stringappend_5730)) = ('''')) then True else undefined)) )))) ))) - else if (((((string_startswith stringappend_5410 (''jal''))) \<and> ((let stringappend_5570 = (string_drop stringappend_5410 ((string_length (''jal'')))) in - if ((case ((spc_matches_prefix stringappend_5570)) of - Some (stringappend_5580,stringappend_5590) => - (let stringappend_5600 = (string_drop stringappend_5570 stringappend_5590) in - if ((case ((reg_name_matches_prefix stringappend_5600 + else if (((((string_startswith stringappend_5590 (''jal''))) \<and> ((let stringappend_5750 = (string_drop stringappend_5590 ((string_length (''jal'')))) in + if ((case ((spc_matches_prefix stringappend_5750)) of + Some (stringappend_5760,stringappend_5770) => + (let stringappend_5780 = (string_drop stringappend_5750 stringappend_5770) in + if ((case ((reg_name_matches_prefix stringappend_5780 :: (( 5 Word.word * ii))option)) of - Some (stringappend_5610,stringappend_5620) => - (let stringappend_5630 = (string_drop stringappend_5600 stringappend_5620) in - if ((case ((sep_matches_prefix stringappend_5630)) of - Some (stringappend_5640,stringappend_5650) => - (let stringappend_5660 = - (string_drop stringappend_5630 stringappend_5650) in + Some (stringappend_5790,stringappend_5800) => + (let stringappend_5810 = (string_drop stringappend_5780 stringappend_5800) in + if ((case ((sep_matches_prefix stringappend_5810)) of + Some (stringappend_5820,stringappend_5830) => + (let stringappend_5840 = + (string_drop stringappend_5810 stringappend_5830) in if ((case ((hex_bits_21_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5660 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5840 :: (( 21 Word.word * ii))option)) of - Some (stringappend_5670,stringappend_5680) => - if(((string_drop stringappend_5660 stringappend_5680)) = ('''')) then + Some (stringappend_5850,stringappend_5860) => + if(((string_drop stringappend_5840 stringappend_5860)) = ('''')) then True else False | None => False )) then @@ -17138,60 +17306,60 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_5570 = (string_drop stringappend_5410 ((string_length (''jal'')))) in + (let stringappend_5750 = (string_drop stringappend_5590 ((string_length (''jal'')))) in (case - (case ((spc_matches_prefix stringappend_5570)) of - Some (stringappend_5580,stringappend_5590) => (stringappend_5580, stringappend_5590) + (case ((spc_matches_prefix stringappend_5750)) of + Some (stringappend_5760,stringappend_5770) => (stringappend_5760, stringappend_5770) ) of - (_, stringappend_5590) => - (let stringappend_5600 = (string_drop stringappend_5570 stringappend_5590) in - (let (rd, stringappend_5620) = - ((case ((reg_name_matches_prefix stringappend_5600 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_5610,stringappend_5620) => (stringappend_5610, stringappend_5620) + (_, stringappend_5770) => + (let stringappend_5780 = (string_drop stringappend_5750 stringappend_5770) in + (let (rd, stringappend_5800) = + ((case ((reg_name_matches_prefix stringappend_5780 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5790,stringappend_5800) => (stringappend_5790, stringappend_5800) )) in - (let stringappend_5630 = (string_drop stringappend_5600 stringappend_5620) in + (let stringappend_5810 = (string_drop stringappend_5780 stringappend_5800) in (case - (case ((sep_matches_prefix stringappend_5630)) of - Some (stringappend_5640,stringappend_5650) => (stringappend_5640, stringappend_5650) + (case ((sep_matches_prefix stringappend_5810)) of + Some (stringappend_5820,stringappend_5830) => (stringappend_5820, stringappend_5830) ) of - (_, stringappend_5650) => - (let stringappend_5660 = (string_drop stringappend_5630 stringappend_5650) in - (let (imm, stringappend_5680) = + (_, stringappend_5830) => + (let stringappend_5840 = (string_drop stringappend_5810 stringappend_5830) in + (let (imm, stringappend_5860) = ((case ((hex_bits_21_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_5660 :: (( 21 Word.word * ii)) option)) of - Some (stringappend_5670,stringappend_5680) => (stringappend_5670, stringappend_5680) + stringappend_5840 :: (( 21 Word.word * ii)) option)) of + Some (stringappend_5850,stringappend_5860) => (stringappend_5850, stringappend_5860) )) in - if(((string_drop stringappend_5660 stringappend_5680)) = ('''')) then + if(((string_drop stringappend_5840 stringappend_5860)) = ('''')) then True else undefined)) )))) )) - else if (((((string_startswith stringappend_5410 (''jalr''))) \<and> ((let stringappend_5700 = (string_drop stringappend_5410 ((string_length (''jalr'')))) in - if ((case ((spc_matches_prefix stringappend_5700)) of - Some (stringappend_5710,stringappend_5720) => - (let stringappend_5730 = (string_drop stringappend_5700 stringappend_5720) in - if ((case ((reg_name_matches_prefix stringappend_5730 + else if (((((string_startswith stringappend_5590 (''jalr''))) \<and> ((let stringappend_5880 = (string_drop stringappend_5590 ((string_length (''jalr'')))) in + if ((case ((spc_matches_prefix stringappend_5880)) of + Some (stringappend_5890,stringappend_5900) => + (let stringappend_5910 = (string_drop stringappend_5880 stringappend_5900) in + if ((case ((reg_name_matches_prefix stringappend_5910 :: (( 5 Word.word * ii))option)) of - Some (stringappend_5740,stringappend_5750) => - (let stringappend_5760 = (string_drop stringappend_5730 stringappend_5750) in - if ((case ((sep_matches_prefix stringappend_5760)) of - Some (stringappend_5770,stringappend_5780) => - (let stringappend_5790 = - (string_drop stringappend_5760 stringappend_5780) in - if ((case ((reg_name_matches_prefix stringappend_5790 + Some (stringappend_5920,stringappend_5930) => + (let stringappend_5940 = (string_drop stringappend_5910 stringappend_5930) in + if ((case ((sep_matches_prefix stringappend_5940)) of + Some (stringappend_5950,stringappend_5960) => + (let stringappend_5970 = + (string_drop stringappend_5940 stringappend_5960) in + if ((case ((reg_name_matches_prefix stringappend_5970 :: (( 5 Word.word * ii))option)) of - Some (stringappend_5800,stringappend_5810) => - (let stringappend_5820 = - (string_drop stringappend_5790 stringappend_5810) in - if ((case ((sep_matches_prefix stringappend_5820)) of - Some (stringappend_5830,stringappend_5840) => - (let stringappend_5850 = - (string_drop stringappend_5820 stringappend_5840) in + Some (stringappend_5980,stringappend_5990) => + (let stringappend_6000 = + (string_drop stringappend_5970 stringappend_5990) in + if ((case ((sep_matches_prefix stringappend_6000)) of + Some (stringappend_6010,stringappend_6020) => + (let stringappend_6030 = + (string_drop stringappend_6000 stringappend_6020) in if ((case ((hex_bits_12_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5850 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_6030 :: (( 12 Word.word * ii))option)) of - Some (stringappend_5860,stringappend_5870) => - if(((string_drop stringappend_5850 stringappend_5870)) = ('''')) then + Some (stringappend_6040,stringappend_6050) => + if(((string_drop stringappend_6030 stringappend_6050)) = ('''')) then True else False | None => False )) then @@ -17217,72 +17385,72 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_5700 = (string_drop stringappend_5410 ((string_length (''jalr'')))) in + (let stringappend_5880 = (string_drop stringappend_5590 ((string_length (''jalr'')))) in (case - (case ((spc_matches_prefix stringappend_5700)) of - Some (stringappend_5710,stringappend_5720) => (stringappend_5710, stringappend_5720) + (case ((spc_matches_prefix stringappend_5880)) of + Some (stringappend_5890,stringappend_5900) => (stringappend_5890, stringappend_5900) ) of - (_, stringappend_5720) => - (let stringappend_5730 = (string_drop stringappend_5700 stringappend_5720) in - (let (rd, stringappend_5750) = - ((case ((reg_name_matches_prefix stringappend_5730 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_5740,stringappend_5750) => (stringappend_5740, stringappend_5750) + (_, stringappend_5900) => + (let stringappend_5910 = (string_drop stringappend_5880 stringappend_5900) in + (let (rd, stringappend_5930) = + ((case ((reg_name_matches_prefix stringappend_5910 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5920,stringappend_5930) => (stringappend_5920, stringappend_5930) )) in - (let stringappend_5760 = (string_drop stringappend_5730 stringappend_5750) in + (let stringappend_5940 = (string_drop stringappend_5910 stringappend_5930) in (case - (case ((sep_matches_prefix stringappend_5760)) of - Some (stringappend_5770,stringappend_5780) => (stringappend_5770, stringappend_5780) + (case ((sep_matches_prefix stringappend_5940)) of + Some (stringappend_5950,stringappend_5960) => (stringappend_5950, stringappend_5960) ) of - (_, stringappend_5780) => - (let stringappend_5790 = (string_drop stringappend_5760 stringappend_5780) in - (let (rs1, stringappend_5810) = - ((case ((reg_name_matches_prefix stringappend_5790 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_5800,stringappend_5810) => (stringappend_5800, stringappend_5810) + (_, stringappend_5960) => + (let stringappend_5970 = (string_drop stringappend_5940 stringappend_5960) in + (let (rs1, stringappend_5990) = + ((case ((reg_name_matches_prefix stringappend_5970 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5980,stringappend_5990) => (stringappend_5980, stringappend_5990) )) in - (let stringappend_5820 = (string_drop stringappend_5790 stringappend_5810) in + (let stringappend_6000 = (string_drop stringappend_5970 stringappend_5990) in (case - (case ((sep_matches_prefix stringappend_5820)) of - Some (stringappend_5830,stringappend_5840) => (stringappend_5830, stringappend_5840) + (case ((sep_matches_prefix stringappend_6000)) of + Some (stringappend_6010,stringappend_6020) => (stringappend_6010, stringappend_6020) ) of - (_, stringappend_5840) => - (let stringappend_5850 = (string_drop stringappend_5820 stringappend_5840) in - (let (imm, stringappend_5870) = + (_, stringappend_6020) => + (let stringappend_6030 = (string_drop stringappend_6000 stringappend_6020) in + (let (imm, stringappend_6050) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_5850 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_5860,stringappend_5870) => (stringappend_5860, stringappend_5870) + stringappend_6030 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_6040,stringappend_6050) => (stringappend_6040, stringappend_6050) )) in - if(((string_drop stringappend_5850 stringappend_5870)) = ('''')) then + if(((string_drop stringappend_6030 stringappend_6050)) = ('''')) then True else undefined)) )))) )))) )) - else if ((case ((btype_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_5890,stringappend_5900) => - (let stringappend_5910 = (string_drop stringappend_5410 stringappend_5900) in - if ((case ((spc_matches_prefix stringappend_5910)) of - Some (stringappend_5920,stringappend_5930) => - (let stringappend_5940 = (string_drop stringappend_5910 stringappend_5930) in - if ((case ((reg_name_matches_prefix stringappend_5940 :: (( 5 Word.word * ii))option)) of - Some (stringappend_5950,stringappend_5960) => - (let stringappend_5970 = (string_drop stringappend_5940 stringappend_5960) in - if ((case ((sep_matches_prefix stringappend_5970)) of - Some (stringappend_5980,stringappend_5990) => - (let stringappend_6000 = (string_drop stringappend_5970 stringappend_5990) in - if ((case ((reg_name_matches_prefix stringappend_6000 + else if ((case ((btype_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_6070,stringappend_6080) => + (let stringappend_6090 = (string_drop stringappend_5590 stringappend_6080) in + if ((case ((spc_matches_prefix stringappend_6090)) of + Some (stringappend_6100,stringappend_6110) => + (let stringappend_6120 = (string_drop stringappend_6090 stringappend_6110) in + if ((case ((reg_name_matches_prefix stringappend_6120 :: (( 5 Word.word * ii))option)) of + Some (stringappend_6130,stringappend_6140) => + (let stringappend_6150 = (string_drop stringappend_6120 stringappend_6140) in + if ((case ((sep_matches_prefix stringappend_6150)) of + Some (stringappend_6160,stringappend_6170) => + (let stringappend_6180 = (string_drop stringappend_6150 stringappend_6170) in + if ((case ((reg_name_matches_prefix stringappend_6180 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6010,stringappend_6020) => - (let stringappend_6030 = - (string_drop stringappend_6000 stringappend_6020) in - if ((case ((sep_matches_prefix stringappend_6030)) of - Some (stringappend_6040,stringappend_6050) => - (let stringappend_6060 = - (string_drop stringappend_6030 stringappend_6050) in + Some (stringappend_6190,stringappend_6200) => + (let stringappend_6210 = + (string_drop stringappend_6180 stringappend_6200) in + if ((case ((sep_matches_prefix stringappend_6210)) of + Some (stringappend_6220,stringappend_6230) => + (let stringappend_6240 = + (string_drop stringappend_6210 stringappend_6230) in if ((case ((hex_bits_13_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_6060 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_6240 :: (( 13 Word.word * ii))option)) of - Some (stringappend_6070,stringappend_6080) => - if(((string_drop stringappend_6060 stringappend_6080)) = ('''')) then + Some (stringappend_6250,stringappend_6260) => + if(((string_drop stringappend_6240 stringappend_6260)) = ('''')) then True else False | None => False )) then @@ -17310,76 +17478,76 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let (op1, stringappend_5900) = - ((case ((btype_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_5890,stringappend_5900) => (stringappend_5890, stringappend_5900) + (let (op1, stringappend_6080) = + ((case ((btype_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_6070,stringappend_6080) => (stringappend_6070, stringappend_6080) )) in - (let stringappend_5910 = (string_drop stringappend_5410 stringappend_5900) in + (let stringappend_6090 = (string_drop stringappend_5590 stringappend_6080) in (case - (case ((spc_matches_prefix stringappend_5910)) of - Some (stringappend_5920,stringappend_5930) => (stringappend_5920, stringappend_5930) + (case ((spc_matches_prefix stringappend_6090)) of + Some (stringappend_6100,stringappend_6110) => (stringappend_6100, stringappend_6110) ) of - (_, stringappend_5930) => - (let stringappend_5940 = (string_drop stringappend_5910 stringappend_5930) in - (let (rs1, stringappend_5960) = - ((case ((reg_name_matches_prefix stringappend_5940 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_5950,stringappend_5960) => (stringappend_5950, stringappend_5960) + (_, stringappend_6110) => + (let stringappend_6120 = (string_drop stringappend_6090 stringappend_6110) in + (let (rs1, stringappend_6140) = + ((case ((reg_name_matches_prefix stringappend_6120 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6130,stringappend_6140) => (stringappend_6130, stringappend_6140) )) in - (let stringappend_5970 = (string_drop stringappend_5940 stringappend_5960) in + (let stringappend_6150 = (string_drop stringappend_6120 stringappend_6140) in (case - (case ((sep_matches_prefix stringappend_5970)) of - Some (stringappend_5980,stringappend_5990) => (stringappend_5980, stringappend_5990) + (case ((sep_matches_prefix stringappend_6150)) of + Some (stringappend_6160,stringappend_6170) => (stringappend_6160, stringappend_6170) ) of - (_, stringappend_5990) => - (let stringappend_6000 = (string_drop stringappend_5970 stringappend_5990) in - (let (rs2, stringappend_6020) = - ((case ((reg_name_matches_prefix stringappend_6000 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6010,stringappend_6020) => (stringappend_6010, stringappend_6020) + (_, stringappend_6170) => + (let stringappend_6180 = (string_drop stringappend_6150 stringappend_6170) in + (let (rs2, stringappend_6200) = + ((case ((reg_name_matches_prefix stringappend_6180 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6190,stringappend_6200) => (stringappend_6190, stringappend_6200) )) in - (let stringappend_6030 = (string_drop stringappend_6000 stringappend_6020) in + (let stringappend_6210 = (string_drop stringappend_6180 stringappend_6200) in (case - (case ((sep_matches_prefix stringappend_6030)) of - Some (stringappend_6040,stringappend_6050) => (stringappend_6040, stringappend_6050) + (case ((sep_matches_prefix stringappend_6210)) of + Some (stringappend_6220,stringappend_6230) => (stringappend_6220, stringappend_6230) ) of - (_, stringappend_6050) => - (let stringappend_6060 = (string_drop stringappend_6030 stringappend_6050) in - (let (imm, stringappend_6080) = + (_, stringappend_6230) => + (let stringappend_6240 = (string_drop stringappend_6210 stringappend_6230) in + (let (imm, stringappend_6260) = ((case ((hex_bits_13_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_6060 :: (( 13 Word.word * ii)) option)) of - Some (stringappend_6070,stringappend_6080) => (stringappend_6070, stringappend_6080) + stringappend_6240 :: (( 13 Word.word * ii)) option)) of + Some (stringappend_6250,stringappend_6260) => (stringappend_6250, stringappend_6260) )) in - if(((string_drop stringappend_6060 stringappend_6080)) = ('''')) then + if(((string_drop stringappend_6240 stringappend_6260)) = ('''')) then True else undefined)) )))) )))) ))) - else if ((case ((itype_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_6100,stringappend_6110) => - (let stringappend_6120 = (string_drop stringappend_5410 stringappend_6110) in - if ((case ((spc_matches_prefix stringappend_6120)) of - Some (stringappend_6130,stringappend_6140) => - (let stringappend_6150 = (string_drop stringappend_6120 stringappend_6140) in - if ((case ((reg_name_matches_prefix stringappend_6150 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6160,stringappend_6170) => - (let stringappend_6180 = (string_drop stringappend_6150 stringappend_6170) in - if ((case ((sep_matches_prefix stringappend_6180)) of - Some (stringappend_6190,stringappend_6200) => - (let stringappend_6210 = (string_drop stringappend_6180 stringappend_6200) in - if ((case ((reg_name_matches_prefix stringappend_6210 + else if ((case ((itype_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_6280,stringappend_6290) => + (let stringappend_6300 = (string_drop stringappend_5590 stringappend_6290) in + if ((case ((spc_matches_prefix stringappend_6300)) of + Some (stringappend_6310,stringappend_6320) => + (let stringappend_6330 = (string_drop stringappend_6300 stringappend_6320) in + if ((case ((reg_name_matches_prefix stringappend_6330 :: (( 5 Word.word * ii))option)) of + Some (stringappend_6340,stringappend_6350) => + (let stringappend_6360 = (string_drop stringappend_6330 stringappend_6350) in + if ((case ((sep_matches_prefix stringappend_6360)) of + Some (stringappend_6370,stringappend_6380) => + (let stringappend_6390 = (string_drop stringappend_6360 stringappend_6380) in + if ((case ((reg_name_matches_prefix stringappend_6390 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6220,stringappend_6230) => - (let stringappend_6240 = - (string_drop stringappend_6210 stringappend_6230) in - if ((case ((sep_matches_prefix stringappend_6240)) of - Some (stringappend_6250,stringappend_6260) => - (let stringappend_6270 = - (string_drop stringappend_6240 stringappend_6260) in + Some (stringappend_6400,stringappend_6410) => + (let stringappend_6420 = + (string_drop stringappend_6390 stringappend_6410) in + if ((case ((sep_matches_prefix stringappend_6420)) of + Some (stringappend_6430,stringappend_6440) => + (let stringappend_6450 = + (string_drop stringappend_6420 stringappend_6440) in if ((case ((hex_bits_12_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_6270 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_6450 :: (( 12 Word.word * ii))option)) of - Some (stringappend_6280,stringappend_6290) => - if(((string_drop stringappend_6270 stringappend_6290)) = ('''')) then + Some (stringappend_6460,stringappend_6470) => + if(((string_drop stringappend_6450 stringappend_6470)) = ('''')) then True else False | None => False )) then @@ -17407,72 +17575,72 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let (op1, stringappend_6110) = - ((case ((itype_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_6100,stringappend_6110) => (stringappend_6100, stringappend_6110) + (let (op1, stringappend_6290) = + ((case ((itype_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_6280,stringappend_6290) => (stringappend_6280, stringappend_6290) )) in - (let stringappend_6120 = (string_drop stringappend_5410 stringappend_6110) in + (let stringappend_6300 = (string_drop stringappend_5590 stringappend_6290) in (case - (case ((spc_matches_prefix stringappend_6120)) of - Some (stringappend_6130,stringappend_6140) => (stringappend_6130, stringappend_6140) + (case ((spc_matches_prefix stringappend_6300)) of + Some (stringappend_6310,stringappend_6320) => (stringappend_6310, stringappend_6320) ) of - (_, stringappend_6140) => - (let stringappend_6150 = (string_drop stringappend_6120 stringappend_6140) in - (let (rd, stringappend_6170) = - ((case ((reg_name_matches_prefix stringappend_6150 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6160,stringappend_6170) => (stringappend_6160, stringappend_6170) + (_, stringappend_6320) => + (let stringappend_6330 = (string_drop stringappend_6300 stringappend_6320) in + (let (rd, stringappend_6350) = + ((case ((reg_name_matches_prefix stringappend_6330 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6340,stringappend_6350) => (stringappend_6340, stringappend_6350) )) in - (let stringappend_6180 = (string_drop stringappend_6150 stringappend_6170) in + (let stringappend_6360 = (string_drop stringappend_6330 stringappend_6350) in (case - (case ((sep_matches_prefix stringappend_6180)) of - Some (stringappend_6190,stringappend_6200) => (stringappend_6190, stringappend_6200) + (case ((sep_matches_prefix stringappend_6360)) of + Some (stringappend_6370,stringappend_6380) => (stringappend_6370, stringappend_6380) ) of - (_, stringappend_6200) => - (let stringappend_6210 = (string_drop stringappend_6180 stringappend_6200) in - (let (rs1, stringappend_6230) = - ((case ((reg_name_matches_prefix stringappend_6210 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6220,stringappend_6230) => (stringappend_6220, stringappend_6230) + (_, stringappend_6380) => + (let stringappend_6390 = (string_drop stringappend_6360 stringappend_6380) in + (let (rs1, stringappend_6410) = + ((case ((reg_name_matches_prefix stringappend_6390 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6400,stringappend_6410) => (stringappend_6400, stringappend_6410) )) in - (let stringappend_6240 = (string_drop stringappend_6210 stringappend_6230) in + (let stringappend_6420 = (string_drop stringappend_6390 stringappend_6410) in (case - (case ((sep_matches_prefix stringappend_6240)) of - Some (stringappend_6250,stringappend_6260) => (stringappend_6250, stringappend_6260) + (case ((sep_matches_prefix stringappend_6420)) of + Some (stringappend_6430,stringappend_6440) => (stringappend_6430, stringappend_6440) ) of - (_, stringappend_6260) => - (let stringappend_6270 = (string_drop stringappend_6240 stringappend_6260) in - (let (imm, stringappend_6290) = + (_, stringappend_6440) => + (let stringappend_6450 = (string_drop stringappend_6420 stringappend_6440) in + (let (imm, stringappend_6470) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_6270 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_6280,stringappend_6290) => (stringappend_6280, stringappend_6290) + stringappend_6450 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_6460,stringappend_6470) => (stringappend_6460, stringappend_6470) )) in - if(((string_drop stringappend_6270 stringappend_6290)) = ('''')) then + if(((string_drop stringappend_6450 stringappend_6470)) = ('''')) then True else undefined)) )))) )))) ))) - else if ((case ((shiftiop_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_6310,stringappend_6320) => - (let stringappend_6330 = (string_drop stringappend_5410 stringappend_6320) in - if ((case ((spc_matches_prefix stringappend_6330)) of - Some (stringappend_6340,stringappend_6350) => - (let stringappend_6360 = (string_drop stringappend_6330 stringappend_6350) in - if ((case ((reg_name_matches_prefix stringappend_6360 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6370,stringappend_6380) => - (let stringappend_6390 = (string_drop stringappend_6360 stringappend_6380) in - if ((case ((sep_matches_prefix stringappend_6390)) of - Some (stringappend_6400,stringappend_6410) => - (let stringappend_6420 = (string_drop stringappend_6390 stringappend_6410) in - if ((case ((reg_name_matches_prefix stringappend_6420 + else if ((case ((shiftiop_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_6490,stringappend_6500) => + (let stringappend_6510 = (string_drop stringappend_5590 stringappend_6500) in + if ((case ((spc_matches_prefix stringappend_6510)) of + Some (stringappend_6520,stringappend_6530) => + (let stringappend_6540 = (string_drop stringappend_6510 stringappend_6530) in + if ((case ((reg_name_matches_prefix stringappend_6540 :: (( 5 Word.word * ii))option)) of + Some (stringappend_6550,stringappend_6560) => + (let stringappend_6570 = (string_drop stringappend_6540 stringappend_6560) in + if ((case ((sep_matches_prefix stringappend_6570)) of + Some (stringappend_6580,stringappend_6590) => + (let stringappend_6600 = (string_drop stringappend_6570 stringappend_6590) in + if ((case ((reg_name_matches_prefix stringappend_6600 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6430,stringappend_6440) => - (let stringappend_6450 = - (string_drop stringappend_6420 stringappend_6440) in + Some (stringappend_6610,stringappend_6620) => + (let stringappend_6630 = + (string_drop stringappend_6600 stringappend_6620) in if ((case ((hex_bits_6_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_6450 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_6630 :: (( 6 Word.word * ii))option)) of - Some (stringappend_6460,stringappend_6470) => - if(((string_drop stringappend_6450 stringappend_6470)) = ('''')) then + Some (stringappend_6640,stringappend_6650) => + if(((string_drop stringappend_6630 stringappend_6650)) = ('''')) then True else False | None => False )) then @@ -17496,68 +17664,68 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let (op1, stringappend_6320) = - ((case ((shiftiop_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_6310,stringappend_6320) => (stringappend_6310, stringappend_6320) + (let (op1, stringappend_6500) = + ((case ((shiftiop_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_6490,stringappend_6500) => (stringappend_6490, stringappend_6500) )) in - (let stringappend_6330 = (string_drop stringappend_5410 stringappend_6320) in + (let stringappend_6510 = (string_drop stringappend_5590 stringappend_6500) in (case - (case ((spc_matches_prefix stringappend_6330)) of - Some (stringappend_6340,stringappend_6350) => (stringappend_6340, stringappend_6350) + (case ((spc_matches_prefix stringappend_6510)) of + Some (stringappend_6520,stringappend_6530) => (stringappend_6520, stringappend_6530) ) of - (_, stringappend_6350) => - (let stringappend_6360 = (string_drop stringappend_6330 stringappend_6350) in - (let (rd, stringappend_6380) = - ((case ((reg_name_matches_prefix stringappend_6360 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6370,stringappend_6380) => (stringappend_6370, stringappend_6380) + (_, stringappend_6530) => + (let stringappend_6540 = (string_drop stringappend_6510 stringappend_6530) in + (let (rd, stringappend_6560) = + ((case ((reg_name_matches_prefix stringappend_6540 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6550,stringappend_6560) => (stringappend_6550, stringappend_6560) )) in - (let stringappend_6390 = (string_drop stringappend_6360 stringappend_6380) in + (let stringappend_6570 = (string_drop stringappend_6540 stringappend_6560) in (case - (case ((sep_matches_prefix stringappend_6390)) of - Some (stringappend_6400,stringappend_6410) => (stringappend_6400, stringappend_6410) + (case ((sep_matches_prefix stringappend_6570)) of + Some (stringappend_6580,stringappend_6590) => (stringappend_6580, stringappend_6590) ) of - (_, stringappend_6410) => - (let stringappend_6420 = (string_drop stringappend_6390 stringappend_6410) in - (let (rs1, stringappend_6440) = - ((case ((reg_name_matches_prefix stringappend_6420 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6430,stringappend_6440) => (stringappend_6430, stringappend_6440) + (_, stringappend_6590) => + (let stringappend_6600 = (string_drop stringappend_6570 stringappend_6590) in + (let (rs1, stringappend_6620) = + ((case ((reg_name_matches_prefix stringappend_6600 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6610,stringappend_6620) => (stringappend_6610, stringappend_6620) )) in - (let stringappend_6450 = (string_drop stringappend_6420 stringappend_6440) in - (let (shamt, stringappend_6470) = + (let stringappend_6630 = (string_drop stringappend_6600 stringappend_6620) in + (let (shamt, stringappend_6650) = ((case ((hex_bits_6_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_6450 :: (( 6 Word.word * ii)) option)) of - Some (stringappend_6460,stringappend_6470) => (stringappend_6460, stringappend_6470) + stringappend_6630 :: (( 6 Word.word * ii)) option)) of + Some (stringappend_6640,stringappend_6650) => (stringappend_6640, stringappend_6650) )) in - if(((string_drop stringappend_6450 stringappend_6470)) = ('''')) then + if(((string_drop stringappend_6630 stringappend_6650)) = ('''')) then True else undefined)))) )))) ))) - else if ((case ((rtype_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_6490,stringappend_6500) => - (let stringappend_6510 = (string_drop stringappend_5410 stringappend_6500) in - if ((case ((spc_matches_prefix stringappend_6510)) of - Some (stringappend_6520,stringappend_6530) => - (let stringappend_6540 = (string_drop stringappend_6510 stringappend_6530) in - if ((case ((reg_name_matches_prefix stringappend_6540 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6550,stringappend_6560) => - (let stringappend_6570 = (string_drop stringappend_6540 stringappend_6560) in - if ((case ((sep_matches_prefix stringappend_6570)) of - Some (stringappend_6580,stringappend_6590) => - (let stringappend_6600 = (string_drop stringappend_6570 stringappend_6590) in - if ((case ((reg_name_matches_prefix stringappend_6600 + else if ((case ((rtype_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_6670,stringappend_6680) => + (let stringappend_6690 = (string_drop stringappend_5590 stringappend_6680) in + if ((case ((spc_matches_prefix stringappend_6690)) of + Some (stringappend_6700,stringappend_6710) => + (let stringappend_6720 = (string_drop stringappend_6690 stringappend_6710) in + if ((case ((reg_name_matches_prefix stringappend_6720 :: (( 5 Word.word * ii))option)) of + Some (stringappend_6730,stringappend_6740) => + (let stringappend_6750 = (string_drop stringappend_6720 stringappend_6740) in + if ((case ((sep_matches_prefix stringappend_6750)) of + Some (stringappend_6760,stringappend_6770) => + (let stringappend_6780 = (string_drop stringappend_6750 stringappend_6770) in + if ((case ((reg_name_matches_prefix stringappend_6780 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6610,stringappend_6620) => - (let stringappend_6630 = - (string_drop stringappend_6600 stringappend_6620) in - if ((case ((sep_matches_prefix stringappend_6630)) of - Some (stringappend_6640,stringappend_6650) => - (let stringappend_6660 = - (string_drop stringappend_6630 stringappend_6650) in - if ((case ((reg_name_matches_prefix stringappend_6660 + Some (stringappend_6790,stringappend_6800) => + (let stringappend_6810 = + (string_drop stringappend_6780 stringappend_6800) in + if ((case ((sep_matches_prefix stringappend_6810)) of + Some (stringappend_6820,stringappend_6830) => + (let stringappend_6840 = + (string_drop stringappend_6810 stringappend_6830) in + if ((case ((reg_name_matches_prefix stringappend_6840 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6670,stringappend_6680) => - if(((string_drop stringappend_6660 stringappend_6680)) = ('''')) then + Some (stringappend_6850,stringappend_6860) => + if(((string_drop stringappend_6840 stringappend_6860)) = ('''')) then True else False | None => False )) then @@ -17585,98 +17753,98 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let (op1, stringappend_6500) = - ((case ((rtype_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_6490,stringappend_6500) => (stringappend_6490, stringappend_6500) + (let (op1, stringappend_6680) = + ((case ((rtype_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_6670,stringappend_6680) => (stringappend_6670, stringappend_6680) )) in - (let stringappend_6510 = (string_drop stringappend_5410 stringappend_6500) in + (let stringappend_6690 = (string_drop stringappend_5590 stringappend_6680) in (case - (case ((spc_matches_prefix stringappend_6510)) of - Some (stringappend_6520,stringappend_6530) => (stringappend_6520, stringappend_6530) + (case ((spc_matches_prefix stringappend_6690)) of + Some (stringappend_6700,stringappend_6710) => (stringappend_6700, stringappend_6710) ) of - (_, stringappend_6530) => - (let stringappend_6540 = (string_drop stringappend_6510 stringappend_6530) in - (let (rd, stringappend_6560) = - ((case ((reg_name_matches_prefix stringappend_6540 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6550,stringappend_6560) => (stringappend_6550, stringappend_6560) + (_, stringappend_6710) => + (let stringappend_6720 = (string_drop stringappend_6690 stringappend_6710) in + (let (rd, stringappend_6740) = + ((case ((reg_name_matches_prefix stringappend_6720 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6730,stringappend_6740) => (stringappend_6730, stringappend_6740) )) in - (let stringappend_6570 = (string_drop stringappend_6540 stringappend_6560) in + (let stringappend_6750 = (string_drop stringappend_6720 stringappend_6740) in (case - (case ((sep_matches_prefix stringappend_6570)) of - Some (stringappend_6580,stringappend_6590) => (stringappend_6580, stringappend_6590) + (case ((sep_matches_prefix stringappend_6750)) of + Some (stringappend_6760,stringappend_6770) => (stringappend_6760, stringappend_6770) ) of - (_, stringappend_6590) => - (let stringappend_6600 = (string_drop stringappend_6570 stringappend_6590) in - (let (rs1, stringappend_6620) = - ((case ((reg_name_matches_prefix stringappend_6600 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6610,stringappend_6620) => (stringappend_6610, stringappend_6620) + (_, stringappend_6770) => + (let stringappend_6780 = (string_drop stringappend_6750 stringappend_6770) in + (let (rs1, stringappend_6800) = + ((case ((reg_name_matches_prefix stringappend_6780 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6790,stringappend_6800) => (stringappend_6790, stringappend_6800) )) in - (let stringappend_6630 = (string_drop stringappend_6600 stringappend_6620) in + (let stringappend_6810 = (string_drop stringappend_6780 stringappend_6800) in (case - (case ((sep_matches_prefix stringappend_6630)) of - Some (stringappend_6640,stringappend_6650) => (stringappend_6640, stringappend_6650) + (case ((sep_matches_prefix stringappend_6810)) of + Some (stringappend_6820,stringappend_6830) => (stringappend_6820, stringappend_6830) ) of - (_, stringappend_6650) => - (let stringappend_6660 = (string_drop stringappend_6630 stringappend_6650) in - (let (rs2, stringappend_6680) = - ((case ((reg_name_matches_prefix stringappend_6660 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6670,stringappend_6680) => (stringappend_6670, stringappend_6680) + (_, stringappend_6830) => + (let stringappend_6840 = (string_drop stringappend_6810 stringappend_6830) in + (let (rs2, stringappend_6860) = + ((case ((reg_name_matches_prefix stringappend_6840 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_6850,stringappend_6860) => (stringappend_6850, stringappend_6860) )) in - if(((string_drop stringappend_6660 stringappend_6680)) = ('''')) then + if(((string_drop stringappend_6840 stringappend_6860)) = ('''')) then True else undefined)) )))) )))) ))) - else if (((((string_startswith stringappend_5410 (''l''))) \<and> ((let stringappend_6700 = (string_drop stringappend_5410 ((string_length (''l'')))) in - if ((case ((size_mnemonic_matches_prefix stringappend_6700)) of - Some (stringappend_6710,stringappend_6720) => - (let stringappend_6730 = (string_drop stringappend_6700 stringappend_6720) in - if ((case ((maybe_u_matches_prefix stringappend_6730)) of - Some (stringappend_6740,stringappend_6750) => - (let stringappend_6760 = (string_drop stringappend_6730 stringappend_6750) in - if ((case ((maybe_aq_matches_prefix stringappend_6760)) of - Some (stringappend_6770,stringappend_6780) => - (let stringappend_6790 = - (string_drop stringappend_6760 stringappend_6780) in - if ((case ((maybe_rl_matches_prefix stringappend_6790)) of - Some (stringappend_6800,stringappend_6810) => - (let stringappend_6820 = - (string_drop stringappend_6790 stringappend_6810) in - if ((case ((spc_matches_prefix stringappend_6820)) of - Some (stringappend_6830,stringappend_6840) => - (let stringappend_6850 = - (string_drop stringappend_6820 stringappend_6840) in - if ((case ((reg_name_matches_prefix stringappend_6850 + else if (((((string_startswith stringappend_5590 (''l''))) \<and> ((let stringappend_6880 = (string_drop stringappend_5590 ((string_length (''l'')))) in + if ((case ((size_mnemonic_matches_prefix stringappend_6880)) of + Some (stringappend_6890,stringappend_6900) => + (let stringappend_6910 = (string_drop stringappend_6880 stringappend_6900) in + if ((case ((maybe_u_matches_prefix stringappend_6910)) of + Some (stringappend_6920,stringappend_6930) => + (let stringappend_6940 = (string_drop stringappend_6910 stringappend_6930) in + if ((case ((maybe_aq_matches_prefix stringappend_6940)) of + Some (stringappend_6950,stringappend_6960) => + (let stringappend_6970 = + (string_drop stringappend_6940 stringappend_6960) in + if ((case ((maybe_rl_matches_prefix stringappend_6970)) of + Some (stringappend_6980,stringappend_6990) => + (let stringappend_7000 = + (string_drop stringappend_6970 stringappend_6990) in + if ((case ((spc_matches_prefix stringappend_7000)) of + Some (stringappend_7010,stringappend_7020) => + (let stringappend_7030 = + (string_drop stringappend_7000 stringappend_7020) in + if ((case ((reg_name_matches_prefix stringappend_7030 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6860,stringappend_6870) => - (let stringappend_6880 = - (string_drop stringappend_6850 stringappend_6870) in - if ((case ((sep_matches_prefix stringappend_6880)) of - Some (stringappend_6890,stringappend_6900) => - (let stringappend_6910 = - (string_drop stringappend_6880 - stringappend_6900) in + Some (stringappend_7040,stringappend_7050) => + (let stringappend_7060 = + (string_drop stringappend_7030 stringappend_7050) in + if ((case ((sep_matches_prefix stringappend_7060)) of + Some (stringappend_7070,stringappend_7080) => + (let stringappend_7090 = + (string_drop stringappend_7060 + stringappend_7080) in if ((case ((reg_name_matches_prefix - stringappend_6910 + stringappend_7090 :: (( 5 Word.word * ii))option)) of - Some (stringappend_6920,stringappend_6930) => - (let stringappend_6940 = - (string_drop stringappend_6910 - stringappend_6930) in + Some (stringappend_7100,stringappend_7110) => + (let stringappend_7120 = + (string_drop stringappend_7090 + stringappend_7110) in if ((case ((sep_matches_prefix - stringappend_6940)) of + stringappend_7120)) of Some - (stringappend_6950,stringappend_6960) => - (let stringappend_6970 = - (string_drop stringappend_6940 - stringappend_6960) in + (stringappend_7130,stringappend_7140) => + (let stringappend_7150 = + (string_drop stringappend_7120 + stringappend_7140) in if ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_6970 + stringappend_7150 :: (( 12 Word.word * ii))option)) of Some - (stringappend_6980,stringappend_6990) => - if(((string_drop stringappend_6970 stringappend_6990)) = ('''')) then + (stringappend_7160,stringappend_7170) => + if(((string_drop stringappend_7150 stringappend_7170)) = ('''')) then True else False | None => False )) then @@ -17718,108 +17886,108 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_6700 = (string_drop stringappend_5410 ((string_length (''l'')))) in - (let (size1, stringappend_6720) = - ((case ((size_mnemonic_matches_prefix stringappend_6700)) of - Some (stringappend_6710,stringappend_6720) => (stringappend_6710, stringappend_6720) + (let stringappend_6880 = (string_drop stringappend_5590 ((string_length (''l'')))) in + (let (size1, stringappend_6900) = + ((case ((size_mnemonic_matches_prefix stringappend_6880)) of + Some (stringappend_6890,stringappend_6900) => (stringappend_6890, stringappend_6900) )) in - (let stringappend_6730 = (string_drop stringappend_6700 stringappend_6720) in - (let (is_unsigned, stringappend_6750) = - ((case ((maybe_u_matches_prefix stringappend_6730)) of - Some (stringappend_6740,stringappend_6750) => (stringappend_6740, stringappend_6750) + (let stringappend_6910 = (string_drop stringappend_6880 stringappend_6900) in + (let (is_unsigned, stringappend_6930) = + ((case ((maybe_u_matches_prefix stringappend_6910)) of + Some (stringappend_6920,stringappend_6930) => (stringappend_6920, stringappend_6930) )) in - (let stringappend_6760 = (string_drop stringappend_6730 stringappend_6750) in - (let (aq, stringappend_6780) = - ((case ((maybe_aq_matches_prefix stringappend_6760)) of - Some (stringappend_6770,stringappend_6780) => (stringappend_6770, stringappend_6780) + (let stringappend_6940 = (string_drop stringappend_6910 stringappend_6930) in + (let (aq, stringappend_6960) = + ((case ((maybe_aq_matches_prefix stringappend_6940)) of + Some (stringappend_6950,stringappend_6960) => (stringappend_6950, stringappend_6960) )) in - (let stringappend_6790 = (string_drop stringappend_6760 stringappend_6780) in - (let (rl, stringappend_6810) = - ((case ((maybe_rl_matches_prefix stringappend_6790)) of - Some (stringappend_6800,stringappend_6810) => (stringappend_6800, stringappend_6810) + (let stringappend_6970 = (string_drop stringappend_6940 stringappend_6960) in + (let (rl, stringappend_6990) = + ((case ((maybe_rl_matches_prefix stringappend_6970)) of + Some (stringappend_6980,stringappend_6990) => (stringappend_6980, stringappend_6990) )) in - (let stringappend_6820 = (string_drop stringappend_6790 stringappend_6810) in + (let stringappend_7000 = (string_drop stringappend_6970 stringappend_6990) in (case - (case ((spc_matches_prefix stringappend_6820)) of - Some (stringappend_6830,stringappend_6840) => (stringappend_6830, stringappend_6840) + (case ((spc_matches_prefix stringappend_7000)) of + Some (stringappend_7010,stringappend_7020) => (stringappend_7010, stringappend_7020) ) of - (_, stringappend_6840) => - (let stringappend_6850 = (string_drop stringappend_6820 stringappend_6840) in - (let (rd, stringappend_6870) = - ((case ((reg_name_matches_prefix stringappend_6850 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6860,stringappend_6870) => (stringappend_6860, stringappend_6870) + (_, stringappend_7020) => + (let stringappend_7030 = (string_drop stringappend_7000 stringappend_7020) in + (let (rd, stringappend_7050) = + ((case ((reg_name_matches_prefix stringappend_7030 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7040,stringappend_7050) => (stringappend_7040, stringappend_7050) )) in - (let stringappend_6880 = (string_drop stringappend_6850 stringappend_6870) in + (let stringappend_7060 = (string_drop stringappend_7030 stringappend_7050) in (case - (case ((sep_matches_prefix stringappend_6880)) of - Some (stringappend_6890,stringappend_6900) => (stringappend_6890, stringappend_6900) + (case ((sep_matches_prefix stringappend_7060)) of + Some (stringappend_7070,stringappend_7080) => (stringappend_7070, stringappend_7080) ) of - (_, stringappend_6900) => - (let stringappend_6910 = (string_drop stringappend_6880 stringappend_6900) in - (let (rs1, stringappend_6930) = - ((case ((reg_name_matches_prefix stringappend_6910 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_6920,stringappend_6930) => (stringappend_6920, stringappend_6930) + (_, stringappend_7080) => + (let stringappend_7090 = (string_drop stringappend_7060 stringappend_7080) in + (let (rs1, stringappend_7110) = + ((case ((reg_name_matches_prefix stringappend_7090 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7100,stringappend_7110) => (stringappend_7100, stringappend_7110) )) in - (let stringappend_6940 = (string_drop stringappend_6910 stringappend_6930) in + (let stringappend_7120 = (string_drop stringappend_7090 stringappend_7110) in (case - (case ((sep_matches_prefix stringappend_6940)) of - Some (stringappend_6950,stringappend_6960) => (stringappend_6950, stringappend_6960) + (case ((sep_matches_prefix stringappend_7120)) of + Some (stringappend_7130,stringappend_7140) => (stringappend_7130, stringappend_7140) ) of - (_, stringappend_6960) => - (let stringappend_6970 = (string_drop stringappend_6940 stringappend_6960) in - (let (imm, stringappend_6990) = + (_, stringappend_7140) => + (let stringappend_7150 = (string_drop stringappend_7120 stringappend_7140) in + (let (imm, stringappend_7170) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_6970 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_6980,stringappend_6990) => (stringappend_6980, stringappend_6990) + stringappend_7150 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_7160,stringappend_7170) => (stringappend_7160, stringappend_7170) )) in - if(((string_drop stringappend_6970 stringappend_6990)) = ('''')) then + if(((string_drop stringappend_7150 stringappend_7170)) = ('''')) then True else undefined)) )))) )))) )))))))))) - else if (((((string_startswith stringappend_5410 (''s''))) \<and> ((let stringappend_7010 = (string_drop stringappend_5410 ((string_length (''s'')))) in - if ((case ((size_mnemonic_matches_prefix stringappend_7010)) of - Some (stringappend_7020,stringappend_7030) => - (let stringappend_7040 = (string_drop stringappend_7010 stringappend_7030) in - if ((case ((maybe_aq_matches_prefix stringappend_7040)) of - Some (stringappend_7050,stringappend_7060) => - (let stringappend_7070 = (string_drop stringappend_7040 stringappend_7060) in - if ((case ((maybe_rl_matches_prefix stringappend_7070)) of - Some (stringappend_7080,stringappend_7090) => - (let stringappend_7100 = - (string_drop stringappend_7070 stringappend_7090) in - if ((case ((spc_matches_prefix stringappend_7100)) of - Some (stringappend_7110,stringappend_7120) => - (let stringappend_7130 = - (string_drop stringappend_7100 stringappend_7120) in - if ((case ((reg_name_matches_prefix stringappend_7130 + else if (((((string_startswith stringappend_5590 (''s''))) \<and> ((let stringappend_7190 = (string_drop stringappend_5590 ((string_length (''s'')))) in + if ((case ((size_mnemonic_matches_prefix stringappend_7190)) of + Some (stringappend_7200,stringappend_7210) => + (let stringappend_7220 = (string_drop stringappend_7190 stringappend_7210) in + if ((case ((maybe_aq_matches_prefix stringappend_7220)) of + Some (stringappend_7230,stringappend_7240) => + (let stringappend_7250 = (string_drop stringappend_7220 stringappend_7240) in + if ((case ((maybe_rl_matches_prefix stringappend_7250)) of + Some (stringappend_7260,stringappend_7270) => + (let stringappend_7280 = + (string_drop stringappend_7250 stringappend_7270) in + if ((case ((spc_matches_prefix stringappend_7280)) of + Some (stringappend_7290,stringappend_7300) => + (let stringappend_7310 = + (string_drop stringappend_7280 stringappend_7300) in + if ((case ((reg_name_matches_prefix stringappend_7310 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7140,stringappend_7150) => - (let stringappend_7160 = - (string_drop stringappend_7130 stringappend_7150) in - if ((case ((sep_matches_prefix stringappend_7160)) of - Some (stringappend_7170,stringappend_7180) => - (let stringappend_7190 = - (string_drop stringappend_7160 stringappend_7180) in - if ((case ((reg_name_matches_prefix stringappend_7190 + Some (stringappend_7320,stringappend_7330) => + (let stringappend_7340 = + (string_drop stringappend_7310 stringappend_7330) in + if ((case ((sep_matches_prefix stringappend_7340)) of + Some (stringappend_7350,stringappend_7360) => + (let stringappend_7370 = + (string_drop stringappend_7340 stringappend_7360) in + if ((case ((reg_name_matches_prefix stringappend_7370 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7200,stringappend_7210) => - (let stringappend_7220 = - (string_drop stringappend_7190 - stringappend_7210) in - if ((case ((sep_matches_prefix stringappend_7220)) of - Some (stringappend_7230,stringappend_7240) => - (let stringappend_7250 = - (string_drop stringappend_7220 - stringappend_7240) in + Some (stringappend_7380,stringappend_7390) => + (let stringappend_7400 = + (string_drop stringappend_7370 + stringappend_7390) in + if ((case ((sep_matches_prefix stringappend_7400)) of + Some (stringappend_7410,stringappend_7420) => + (let stringappend_7430 = + (string_drop stringappend_7400 + stringappend_7420) in if ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_7250 + stringappend_7430 :: (( 12 Word.word * ii))option)) of Some - (stringappend_7260,stringappend_7270) => - if(((string_drop stringappend_7250 stringappend_7270)) = ('''')) then + (stringappend_7440,stringappend_7450) => + if(((string_drop stringappend_7430 stringappend_7450)) = ('''')) then True else False | None => False )) then @@ -17857,87 +18025,87 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_7010 = (string_drop stringappend_5410 ((string_length (''s'')))) in - (let (size1, stringappend_7030) = - ((case ((size_mnemonic_matches_prefix stringappend_7010)) of - Some (stringappend_7020,stringappend_7030) => (stringappend_7020, stringappend_7030) + (let stringappend_7190 = (string_drop stringappend_5590 ((string_length (''s'')))) in + (let (size1, stringappend_7210) = + ((case ((size_mnemonic_matches_prefix stringappend_7190)) of + Some (stringappend_7200,stringappend_7210) => (stringappend_7200, stringappend_7210) )) in - (let stringappend_7040 = (string_drop stringappend_7010 stringappend_7030) in - (let (aq, stringappend_7060) = - ((case ((maybe_aq_matches_prefix stringappend_7040)) of - Some (stringappend_7050,stringappend_7060) => (stringappend_7050, stringappend_7060) + (let stringappend_7220 = (string_drop stringappend_7190 stringappend_7210) in + (let (aq, stringappend_7240) = + ((case ((maybe_aq_matches_prefix stringappend_7220)) of + Some (stringappend_7230,stringappend_7240) => (stringappend_7230, stringappend_7240) )) in - (let stringappend_7070 = (string_drop stringappend_7040 stringappend_7060) in - (let (rl, stringappend_7090) = - ((case ((maybe_rl_matches_prefix stringappend_7070)) of - Some (stringappend_7080,stringappend_7090) => (stringappend_7080, stringappend_7090) + (let stringappend_7250 = (string_drop stringappend_7220 stringappend_7240) in + (let (rl, stringappend_7270) = + ((case ((maybe_rl_matches_prefix stringappend_7250)) of + Some (stringappend_7260,stringappend_7270) => (stringappend_7260, stringappend_7270) )) in - (let stringappend_7100 = (string_drop stringappend_7070 stringappend_7090) in + (let stringappend_7280 = (string_drop stringappend_7250 stringappend_7270) in (case - (case ((spc_matches_prefix stringappend_7100)) of - Some (stringappend_7110,stringappend_7120) => (stringappend_7110, stringappend_7120) + (case ((spc_matches_prefix stringappend_7280)) of + Some (stringappend_7290,stringappend_7300) => (stringappend_7290, stringappend_7300) ) of - (_, stringappend_7120) => - (let stringappend_7130 = (string_drop stringappend_7100 stringappend_7120) in - (let (rd, stringappend_7150) = - ((case ((reg_name_matches_prefix stringappend_7130 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7140,stringappend_7150) => (stringappend_7140, stringappend_7150) + (_, stringappend_7300) => + (let stringappend_7310 = (string_drop stringappend_7280 stringappend_7300) in + (let (rd, stringappend_7330) = + ((case ((reg_name_matches_prefix stringappend_7310 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7320,stringappend_7330) => (stringappend_7320, stringappend_7330) )) in - (let stringappend_7160 = (string_drop stringappend_7130 stringappend_7150) in + (let stringappend_7340 = (string_drop stringappend_7310 stringappend_7330) in (case - (case ((sep_matches_prefix stringappend_7160)) of - Some (stringappend_7170,stringappend_7180) => (stringappend_7170, stringappend_7180) + (case ((sep_matches_prefix stringappend_7340)) of + Some (stringappend_7350,stringappend_7360) => (stringappend_7350, stringappend_7360) ) of - (_, stringappend_7180) => - (let stringappend_7190 = (string_drop stringappend_7160 stringappend_7180) in - (let (rs1, stringappend_7210) = - ((case ((reg_name_matches_prefix stringappend_7190 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7200,stringappend_7210) => (stringappend_7200, stringappend_7210) + (_, stringappend_7360) => + (let stringappend_7370 = (string_drop stringappend_7340 stringappend_7360) in + (let (rs1, stringappend_7390) = + ((case ((reg_name_matches_prefix stringappend_7370 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7380,stringappend_7390) => (stringappend_7380, stringappend_7390) )) in - (let stringappend_7220 = (string_drop stringappend_7190 stringappend_7210) in + (let stringappend_7400 = (string_drop stringappend_7370 stringappend_7390) in (case - (case ((sep_matches_prefix stringappend_7220)) of - Some (stringappend_7230,stringappend_7240) => (stringappend_7230, stringappend_7240) + (case ((sep_matches_prefix stringappend_7400)) of + Some (stringappend_7410,stringappend_7420) => (stringappend_7410, stringappend_7420) ) of - (_, stringappend_7240) => - (let stringappend_7250 = (string_drop stringappend_7220 stringappend_7240) in - (let (imm, stringappend_7270) = + (_, stringappend_7420) => + (let stringappend_7430 = (string_drop stringappend_7400 stringappend_7420) in + (let (imm, stringappend_7450) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_7250 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_7260,stringappend_7270) => (stringappend_7260, stringappend_7270) + stringappend_7430 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_7440,stringappend_7450) => (stringappend_7440, stringappend_7450) )) in - if(((string_drop stringappend_7250 stringappend_7270)) = ('''')) then + if(((string_drop stringappend_7430 stringappend_7450)) = ('''')) then True else undefined)) )))) )))) )))))))) - else if (((((string_startswith stringappend_5410 (''addiw''))) \<and> ((let stringappend_7290 = (string_drop stringappend_5410 ((string_length (''addiw'')))) in - if ((case ((spc_matches_prefix stringappend_7290)) of - Some (stringappend_7300,stringappend_7310) => - (let stringappend_7320 = (string_drop stringappend_7290 stringappend_7310) in - if ((case ((reg_name_matches_prefix stringappend_7320 + else if (((((string_startswith stringappend_5590 (''addiw''))) \<and> ((let stringappend_7470 = (string_drop stringappend_5590 ((string_length (''addiw'')))) in + if ((case ((spc_matches_prefix stringappend_7470)) of + Some (stringappend_7480,stringappend_7490) => + (let stringappend_7500 = (string_drop stringappend_7470 stringappend_7490) in + if ((case ((reg_name_matches_prefix stringappend_7500 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7330,stringappend_7340) => - (let stringappend_7350 = (string_drop stringappend_7320 stringappend_7340) in - if ((case ((sep_matches_prefix stringappend_7350)) of - Some (stringappend_7360,stringappend_7370) => - (let stringappend_7380 = - (string_drop stringappend_7350 stringappend_7370) in - if ((case ((reg_name_matches_prefix stringappend_7380 + Some (stringappend_7510,stringappend_7520) => + (let stringappend_7530 = (string_drop stringappend_7500 stringappend_7520) in + if ((case ((sep_matches_prefix stringappend_7530)) of + Some (stringappend_7540,stringappend_7550) => + (let stringappend_7560 = + (string_drop stringappend_7530 stringappend_7550) in + if ((case ((reg_name_matches_prefix stringappend_7560 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7390,stringappend_7400) => - (let stringappend_7410 = - (string_drop stringappend_7380 stringappend_7400) in - if ((case ((sep_matches_prefix stringappend_7410)) of - Some (stringappend_7420,stringappend_7430) => - (let stringappend_7440 = - (string_drop stringappend_7410 stringappend_7430) in + Some (stringappend_7570,stringappend_7580) => + (let stringappend_7590 = + (string_drop stringappend_7560 stringappend_7580) in + if ((case ((sep_matches_prefix stringappend_7590)) of + Some (stringappend_7600,stringappend_7610) => + (let stringappend_7620 = + (string_drop stringappend_7590 stringappend_7610) in if ((case ((hex_bits_12_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_7440 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_7620 :: (( 12 Word.word * ii))option)) of - Some (stringappend_7450,stringappend_7460) => - if(((string_drop stringappend_7440 stringappend_7460)) = ('''')) then + Some (stringappend_7630,stringappend_7640) => + if(((string_drop stringappend_7620 stringappend_7640)) = ('''')) then True else False | None => False )) then @@ -17963,72 +18131,72 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_7290 = (string_drop stringappend_5410 ((string_length (''addiw'')))) in + (let stringappend_7470 = (string_drop stringappend_5590 ((string_length (''addiw'')))) in (case - (case ((spc_matches_prefix stringappend_7290)) of - Some (stringappend_7300,stringappend_7310) => (stringappend_7300, stringappend_7310) + (case ((spc_matches_prefix stringappend_7470)) of + Some (stringappend_7480,stringappend_7490) => (stringappend_7480, stringappend_7490) ) of - (_, stringappend_7310) => - (let stringappend_7320 = (string_drop stringappend_7290 stringappend_7310) in - (let (rd, stringappend_7340) = - ((case ((reg_name_matches_prefix stringappend_7320 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7330,stringappend_7340) => (stringappend_7330, stringappend_7340) + (_, stringappend_7490) => + (let stringappend_7500 = (string_drop stringappend_7470 stringappend_7490) in + (let (rd, stringappend_7520) = + ((case ((reg_name_matches_prefix stringappend_7500 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7510,stringappend_7520) => (stringappend_7510, stringappend_7520) )) in - (let stringappend_7350 = (string_drop stringappend_7320 stringappend_7340) in + (let stringappend_7530 = (string_drop stringappend_7500 stringappend_7520) in (case - (case ((sep_matches_prefix stringappend_7350)) of - Some (stringappend_7360,stringappend_7370) => (stringappend_7360, stringappend_7370) + (case ((sep_matches_prefix stringappend_7530)) of + Some (stringappend_7540,stringappend_7550) => (stringappend_7540, stringappend_7550) ) of - (_, stringappend_7370) => - (let stringappend_7380 = (string_drop stringappend_7350 stringappend_7370) in - (let (rs1, stringappend_7400) = - ((case ((reg_name_matches_prefix stringappend_7380 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7390,stringappend_7400) => (stringappend_7390, stringappend_7400) + (_, stringappend_7550) => + (let stringappend_7560 = (string_drop stringappend_7530 stringappend_7550) in + (let (rs1, stringappend_7580) = + ((case ((reg_name_matches_prefix stringappend_7560 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7570,stringappend_7580) => (stringappend_7570, stringappend_7580) )) in - (let stringappend_7410 = (string_drop stringappend_7380 stringappend_7400) in + (let stringappend_7590 = (string_drop stringappend_7560 stringappend_7580) in (case - (case ((sep_matches_prefix stringappend_7410)) of - Some (stringappend_7420,stringappend_7430) => (stringappend_7420, stringappend_7430) + (case ((sep_matches_prefix stringappend_7590)) of + Some (stringappend_7600,stringappend_7610) => (stringappend_7600, stringappend_7610) ) of - (_, stringappend_7430) => - (let stringappend_7440 = (string_drop stringappend_7410 stringappend_7430) in - (let (imm, stringappend_7460) = + (_, stringappend_7610) => + (let stringappend_7620 = (string_drop stringappend_7590 stringappend_7610) in + (let (imm, stringappend_7640) = ((case ((hex_bits_12_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_7440 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_7450,stringappend_7460) => (stringappend_7450, stringappend_7460) + stringappend_7620 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_7630,stringappend_7640) => (stringappend_7630, stringappend_7640) )) in - if(((string_drop stringappend_7440 stringappend_7460)) = ('''')) then + if(((string_drop stringappend_7620 stringappend_7640)) = ('''')) then True else undefined)) )))) )))) )) - else if ((case ((shiftw_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_7480,stringappend_7490) => - (let stringappend_7500 = (string_drop stringappend_5410 stringappend_7490) in - if ((case ((spc_matches_prefix stringappend_7500)) of - Some (stringappend_7510,stringappend_7520) => - (let stringappend_7530 = (string_drop stringappend_7500 stringappend_7520) in - if ((case ((reg_name_matches_prefix stringappend_7530 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7540,stringappend_7550) => - (let stringappend_7560 = (string_drop stringappend_7530 stringappend_7550) in - if ((case ((sep_matches_prefix stringappend_7560)) of - Some (stringappend_7570,stringappend_7580) => - (let stringappend_7590 = (string_drop stringappend_7560 stringappend_7580) in - if ((case ((reg_name_matches_prefix stringappend_7590 + else if ((case ((shiftw_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_7660,stringappend_7670) => + (let stringappend_7680 = (string_drop stringappend_5590 stringappend_7670) in + if ((case ((spc_matches_prefix stringappend_7680)) of + Some (stringappend_7690,stringappend_7700) => + (let stringappend_7710 = (string_drop stringappend_7680 stringappend_7700) in + if ((case ((reg_name_matches_prefix stringappend_7710 :: (( 5 Word.word * ii))option)) of + Some (stringappend_7720,stringappend_7730) => + (let stringappend_7740 = (string_drop stringappend_7710 stringappend_7730) in + if ((case ((sep_matches_prefix stringappend_7740)) of + Some (stringappend_7750,stringappend_7760) => + (let stringappend_7770 = (string_drop stringappend_7740 stringappend_7760) in + if ((case ((reg_name_matches_prefix stringappend_7770 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7600,stringappend_7610) => - (let stringappend_7620 = - (string_drop stringappend_7590 stringappend_7610) in - if ((case ((sep_matches_prefix stringappend_7620)) of - Some (stringappend_7630,stringappend_7640) => - (let stringappend_7650 = - (string_drop stringappend_7620 stringappend_7640) in + Some (stringappend_7780,stringappend_7790) => + (let stringappend_7800 = + (string_drop stringappend_7770 stringappend_7790) in + if ((case ((sep_matches_prefix stringappend_7800)) of + Some (stringappend_7810,stringappend_7820) => + (let stringappend_7830 = + (string_drop stringappend_7800 stringappend_7820) in if ((case ((hex_bits_5_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_7650 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_7830 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7660,stringappend_7670) => - if(((string_drop stringappend_7650 stringappend_7670)) = ('''')) then + Some (stringappend_7840,stringappend_7850) => + if(((string_drop stringappend_7830 stringappend_7850)) = ('''')) then True else False | None => False )) then @@ -18056,75 +18224,75 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let (op1, stringappend_7490) = - ((case ((shiftw_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_7480,stringappend_7490) => (stringappend_7480, stringappend_7490) + (let (op1, stringappend_7670) = + ((case ((shiftw_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_7660,stringappend_7670) => (stringappend_7660, stringappend_7670) )) in - (let stringappend_7500 = (string_drop stringappend_5410 stringappend_7490) in + (let stringappend_7680 = (string_drop stringappend_5590 stringappend_7670) in (case - (case ((spc_matches_prefix stringappend_7500)) of - Some (stringappend_7510,stringappend_7520) => (stringappend_7510, stringappend_7520) + (case ((spc_matches_prefix stringappend_7680)) of + Some (stringappend_7690,stringappend_7700) => (stringappend_7690, stringappend_7700) ) of - (_, stringappend_7520) => - (let stringappend_7530 = (string_drop stringappend_7500 stringappend_7520) in - (let (rd, stringappend_7550) = - ((case ((reg_name_matches_prefix stringappend_7530 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7540,stringappend_7550) => (stringappend_7540, stringappend_7550) + (_, stringappend_7700) => + (let stringappend_7710 = (string_drop stringappend_7680 stringappend_7700) in + (let (rd, stringappend_7730) = + ((case ((reg_name_matches_prefix stringappend_7710 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7720,stringappend_7730) => (stringappend_7720, stringappend_7730) )) in - (let stringappend_7560 = (string_drop stringappend_7530 stringappend_7550) in + (let stringappend_7740 = (string_drop stringappend_7710 stringappend_7730) in (case - (case ((sep_matches_prefix stringappend_7560)) of - Some (stringappend_7570,stringappend_7580) => (stringappend_7570, stringappend_7580) + (case ((sep_matches_prefix stringappend_7740)) of + Some (stringappend_7750,stringappend_7760) => (stringappend_7750, stringappend_7760) ) of - (_, stringappend_7580) => - (let stringappend_7590 = (string_drop stringappend_7560 stringappend_7580) in - (let (rs1, stringappend_7610) = - ((case ((reg_name_matches_prefix stringappend_7590 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7600,stringappend_7610) => (stringappend_7600, stringappend_7610) + (_, stringappend_7760) => + (let stringappend_7770 = (string_drop stringappend_7740 stringappend_7760) in + (let (rs1, stringappend_7790) = + ((case ((reg_name_matches_prefix stringappend_7770 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7780,stringappend_7790) => (stringappend_7780, stringappend_7790) )) in - (let stringappend_7620 = (string_drop stringappend_7590 stringappend_7610) in + (let stringappend_7800 = (string_drop stringappend_7770 stringappend_7790) in (case - (case ((sep_matches_prefix stringappend_7620)) of - Some (stringappend_7630,stringappend_7640) => (stringappend_7630, stringappend_7640) + (case ((sep_matches_prefix stringappend_7800)) of + Some (stringappend_7810,stringappend_7820) => (stringappend_7810, stringappend_7820) ) of - (_, stringappend_7640) => - (let stringappend_7650 = (string_drop stringappend_7620 stringappend_7640) in - (let (shamt, stringappend_7670) = + (_, stringappend_7820) => + (let stringappend_7830 = (string_drop stringappend_7800 stringappend_7820) in + (let (shamt, stringappend_7850) = ((case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_7650 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7660,stringappend_7670) => (stringappend_7660, stringappend_7670) + stringappend_7830 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7840,stringappend_7850) => (stringappend_7840, stringappend_7850) )) in - if(((string_drop stringappend_7650 stringappend_7670)) = ('''')) then + if(((string_drop stringappend_7830 stringappend_7850)) = ('''')) then True else undefined)) )))) )))) ))) - else if ((case ((rtypew_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_7690,stringappend_7700) => - (let stringappend_7710 = (string_drop stringappend_5410 stringappend_7700) in - if ((case ((spc_matches_prefix stringappend_7710)) of - Some (stringappend_7720,stringappend_7730) => - (let stringappend_7740 = (string_drop stringappend_7710 stringappend_7730) in - if ((case ((reg_name_matches_prefix stringappend_7740 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7750,stringappend_7760) => - (let stringappend_7770 = (string_drop stringappend_7740 stringappend_7760) in - if ((case ((sep_matches_prefix stringappend_7770)) of - Some (stringappend_7780,stringappend_7790) => - (let stringappend_7800 = (string_drop stringappend_7770 stringappend_7790) in - if ((case ((reg_name_matches_prefix stringappend_7800 + else if ((case ((rtypew_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_7870,stringappend_7880) => + (let stringappend_7890 = (string_drop stringappend_5590 stringappend_7880) in + if ((case ((spc_matches_prefix stringappend_7890)) of + Some (stringappend_7900,stringappend_7910) => + (let stringappend_7920 = (string_drop stringappend_7890 stringappend_7910) in + if ((case ((reg_name_matches_prefix stringappend_7920 :: (( 5 Word.word * ii))option)) of + Some (stringappend_7930,stringappend_7940) => + (let stringappend_7950 = (string_drop stringappend_7920 stringappend_7940) in + if ((case ((sep_matches_prefix stringappend_7950)) of + Some (stringappend_7960,stringappend_7970) => + (let stringappend_7980 = (string_drop stringappend_7950 stringappend_7970) in + if ((case ((reg_name_matches_prefix stringappend_7980 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7810,stringappend_7820) => - (let stringappend_7830 = - (string_drop stringappend_7800 stringappend_7820) in - if ((case ((sep_matches_prefix stringappend_7830)) of - Some (stringappend_7840,stringappend_7850) => - (let stringappend_7860 = - (string_drop stringappend_7830 stringappend_7850) in - if ((case ((reg_name_matches_prefix stringappend_7860 + Some (stringappend_7990,stringappend_8000) => + (let stringappend_8010 = + (string_drop stringappend_7980 stringappend_8000) in + if ((case ((sep_matches_prefix stringappend_8010)) of + Some (stringappend_8020,stringappend_8030) => + (let stringappend_8040 = + (string_drop stringappend_8010 stringappend_8030) in + if ((case ((reg_name_matches_prefix stringappend_8040 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7870,stringappend_7880) => - if(((string_drop stringappend_7860 stringappend_7880)) = ('''')) then + Some (stringappend_8050,stringappend_8060) => + if(((string_drop stringappend_8040 stringappend_8060)) = ('''')) then True else False | None => False )) then @@ -18152,73 +18320,155 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let (op1, stringappend_7700) = - ((case ((rtypew_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_7690,stringappend_7700) => (stringappend_7690, stringappend_7700) + (let (op1, stringappend_7880) = + ((case ((rtypew_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_7870,stringappend_7880) => (stringappend_7870, stringappend_7880) )) in - (let stringappend_7710 = (string_drop stringappend_5410 stringappend_7700) in + (let stringappend_7890 = (string_drop stringappend_5590 stringappend_7880) in (case - (case ((spc_matches_prefix stringappend_7710)) of - Some (stringappend_7720,stringappend_7730) => (stringappend_7720, stringappend_7730) + (case ((spc_matches_prefix stringappend_7890)) of + Some (stringappend_7900,stringappend_7910) => (stringappend_7900, stringappend_7910) ) of - (_, stringappend_7730) => - (let stringappend_7740 = (string_drop stringappend_7710 stringappend_7730) in - (let (rd, stringappend_7760) = - ((case ((reg_name_matches_prefix stringappend_7740 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7750,stringappend_7760) => (stringappend_7750, stringappend_7760) + (_, stringappend_7910) => + (let stringappend_7920 = (string_drop stringappend_7890 stringappend_7910) in + (let (rd, stringappend_7940) = + ((case ((reg_name_matches_prefix stringappend_7920 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7930,stringappend_7940) => (stringappend_7930, stringappend_7940) )) in - (let stringappend_7770 = (string_drop stringappend_7740 stringappend_7760) in + (let stringappend_7950 = (string_drop stringappend_7920 stringappend_7940) in (case - (case ((sep_matches_prefix stringappend_7770)) of - Some (stringappend_7780,stringappend_7790) => (stringappend_7780, stringappend_7790) + (case ((sep_matches_prefix stringappend_7950)) of + Some (stringappend_7960,stringappend_7970) => (stringappend_7960, stringappend_7970) ) of - (_, stringappend_7790) => - (let stringappend_7800 = (string_drop stringappend_7770 stringappend_7790) in - (let (rs1, stringappend_7820) = - ((case ((reg_name_matches_prefix stringappend_7800 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7810,stringappend_7820) => (stringappend_7810, stringappend_7820) + (_, stringappend_7970) => + (let stringappend_7980 = (string_drop stringappend_7950 stringappend_7970) in + (let (rs1, stringappend_8000) = + ((case ((reg_name_matches_prefix stringappend_7980 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_7990,stringappend_8000) => (stringappend_7990, stringappend_8000) )) in - (let stringappend_7830 = (string_drop stringappend_7800 stringappend_7820) in + (let stringappend_8010 = (string_drop stringappend_7980 stringappend_8000) in (case - (case ((sep_matches_prefix stringappend_7830)) of - Some (stringappend_7840,stringappend_7850) => (stringappend_7840, stringappend_7850) + (case ((sep_matches_prefix stringappend_8010)) of + Some (stringappend_8020,stringappend_8030) => (stringappend_8020, stringappend_8030) ) of - (_, stringappend_7850) => - (let stringappend_7860 = (string_drop stringappend_7830 stringappend_7850) in - (let (rs2, stringappend_7880) = - ((case ((reg_name_matches_prefix stringappend_7860 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7870,stringappend_7880) => (stringappend_7870, stringappend_7880) + (_, stringappend_8030) => + (let stringappend_8040 = (string_drop stringappend_8010 stringappend_8030) in + (let (rs2, stringappend_8060) = + ((case ((reg_name_matches_prefix stringappend_8040 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8050,stringappend_8060) => (stringappend_8050, stringappend_8060) )) in - if(((string_drop stringappend_7860 stringappend_7880)) = ('''')) then + if(((string_drop stringappend_8040 stringappend_8060)) = ('''')) then True else undefined)) )))) )))) ))) - else if ((case ((mul_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_7900,stringappend_7910) => - (let stringappend_7920 = (string_drop stringappend_5410 stringappend_7910) in - if ((case ((spc_matches_prefix stringappend_7920)) of - Some (stringappend_7930,stringappend_7940) => - (let stringappend_7950 = (string_drop stringappend_7920 stringappend_7940) in - if ((case ((reg_name_matches_prefix stringappend_7950 :: (( 5 Word.word * ii))option)) of - Some (stringappend_7960,stringappend_7970) => - (let stringappend_7980 = (string_drop stringappend_7950 stringappend_7970) in - if ((case ((sep_matches_prefix stringappend_7980)) of - Some (stringappend_7990,stringappend_8000) => - (let stringappend_8010 = (string_drop stringappend_7980 stringappend_8000) in - if ((case ((reg_name_matches_prefix stringappend_8010 + else if ((case ((shiftiwop_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_8080,stringappend_8090) => + (let stringappend_8100 = (string_drop stringappend_5590 stringappend_8090) in + if ((case ((spc_matches_prefix stringappend_8100)) of + Some (stringappend_8110,stringappend_8120) => + (let stringappend_8130 = (string_drop stringappend_8100 stringappend_8120) in + if ((case ((reg_name_matches_prefix stringappend_8130 :: (( 5 Word.word * ii))option)) of + Some (stringappend_8140,stringappend_8150) => + (let stringappend_8160 = (string_drop stringappend_8130 stringappend_8150) in + if ((case ((sep_matches_prefix stringappend_8160)) of + Some (stringappend_8170,stringappend_8180) => + (let stringappend_8190 = (string_drop stringappend_8160 stringappend_8180) in + if ((case ((reg_name_matches_prefix stringappend_8190 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8020,stringappend_8030) => - (let stringappend_8040 = - (string_drop stringappend_8010 stringappend_8030) in - if ((case ((sep_matches_prefix stringappend_8040)) of - Some (stringappend_8050,stringappend_8060) => - (let stringappend_8070 = - (string_drop stringappend_8040 stringappend_8060) in - if ((case ((reg_name_matches_prefix stringappend_8070 + Some (stringappend_8200,stringappend_8210) => + (let stringappend_8220 = + (string_drop stringappend_8190 stringappend_8210) in + if ((case ((hex_bits_5_matches_prefix + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_8220 + :: (( 5 Word.word * ii))option)) of + Some (stringappend_8230,stringappend_8240) => + if(((string_drop stringappend_8220 stringappend_8240)) = ('''')) then + True else False + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + (let (op1, stringappend_8090) = + ((case ((shiftiwop_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_8080,stringappend_8090) => (stringappend_8080, stringappend_8090) + )) in + (let stringappend_8100 = (string_drop stringappend_5590 stringappend_8090) in + (case + (case ((spc_matches_prefix stringappend_8100)) of + Some (stringappend_8110,stringappend_8120) => (stringappend_8110, stringappend_8120) + ) of + (_, stringappend_8120) => + (let stringappend_8130 = (string_drop stringappend_8100 stringappend_8120) in + (let (rd, stringappend_8150) = + ((case ((reg_name_matches_prefix stringappend_8130 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8140,stringappend_8150) => (stringappend_8140, stringappend_8150) + )) in + (let stringappend_8160 = (string_drop stringappend_8130 stringappend_8150) in + (case + (case ((sep_matches_prefix stringappend_8160)) of + Some (stringappend_8170,stringappend_8180) => (stringappend_8170, stringappend_8180) + ) of + (_, stringappend_8180) => + (let stringappend_8190 = (string_drop stringappend_8160 stringappend_8180) in + (let (rs1, stringappend_8210) = + ((case ((reg_name_matches_prefix stringappend_8190 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8200,stringappend_8210) => (stringappend_8200, stringappend_8210) + )) in + (let stringappend_8220 = (string_drop stringappend_8190 stringappend_8210) in + (let (shamt, stringappend_8240) = + ((case ((hex_bits_5_matches_prefix + instance_Sail2_values_Bitvector_Machine_word_mword_dict + stringappend_8220 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8230,stringappend_8240) => (stringappend_8230, stringappend_8240) + )) in + if(((string_drop stringappend_8220 stringappend_8240)) = ('''')) then + True else undefined)))) + )))) + ))) + else if ((case ((mul_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_8260,stringappend_8270) => + (let stringappend_8280 = (string_drop stringappend_5590 stringappend_8270) in + if ((case ((spc_matches_prefix stringappend_8280)) of + Some (stringappend_8290,stringappend_8300) => + (let stringappend_8310 = (string_drop stringappend_8280 stringappend_8300) in + if ((case ((reg_name_matches_prefix stringappend_8310 :: (( 5 Word.word * ii))option)) of + Some (stringappend_8320,stringappend_8330) => + (let stringappend_8340 = (string_drop stringappend_8310 stringappend_8330) in + if ((case ((sep_matches_prefix stringappend_8340)) of + Some (stringappend_8350,stringappend_8360) => + (let stringappend_8370 = (string_drop stringappend_8340 stringappend_8360) in + if ((case ((reg_name_matches_prefix stringappend_8370 + :: (( 5 Word.word * ii))option)) of + Some (stringappend_8380,stringappend_8390) => + (let stringappend_8400 = + (string_drop stringappend_8370 stringappend_8390) in + if ((case ((sep_matches_prefix stringappend_8400)) of + Some (stringappend_8410,stringappend_8420) => + (let stringappend_8430 = + (string_drop stringappend_8400 stringappend_8420) in + if ((case ((reg_name_matches_prefix stringappend_8430 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8080,stringappend_8090) => - if(((string_drop stringappend_8070 stringappend_8090)) = ('''')) then + Some (stringappend_8440,stringappend_8450) => + if(((string_drop stringappend_8430 stringappend_8450)) = ('''')) then True else False | None => False )) then @@ -18246,77 +18496,77 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where else False) | None => False )) then - (let ((high, signed1, signed2), stringappend_7910) = - ((case ((mul_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_7900,stringappend_7910) => (stringappend_7900, stringappend_7910) + (let ((high, signed1, signed2), stringappend_8270) = + ((case ((mul_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_8260,stringappend_8270) => (stringappend_8260, stringappend_8270) )) in - (let stringappend_7920 = (string_drop stringappend_5410 stringappend_7910) in + (let stringappend_8280 = (string_drop stringappend_5590 stringappend_8270) in (case - (case ((spc_matches_prefix stringappend_7920)) of - Some (stringappend_7930,stringappend_7940) => (stringappend_7930, stringappend_7940) + (case ((spc_matches_prefix stringappend_8280)) of + Some (stringappend_8290,stringappend_8300) => (stringappend_8290, stringappend_8300) ) of - (_, stringappend_7940) => - (let stringappend_7950 = (string_drop stringappend_7920 stringappend_7940) in - (let (rd, stringappend_7970) = - ((case ((reg_name_matches_prefix stringappend_7950 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_7960,stringappend_7970) => (stringappend_7960, stringappend_7970) + (_, stringappend_8300) => + (let stringappend_8310 = (string_drop stringappend_8280 stringappend_8300) in + (let (rd, stringappend_8330) = + ((case ((reg_name_matches_prefix stringappend_8310 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8320,stringappend_8330) => (stringappend_8320, stringappend_8330) )) in - (let stringappend_7980 = (string_drop stringappend_7950 stringappend_7970) in + (let stringappend_8340 = (string_drop stringappend_8310 stringappend_8330) in (case - (case ((sep_matches_prefix stringappend_7980)) of - Some (stringappend_7990,stringappend_8000) => (stringappend_7990, stringappend_8000) + (case ((sep_matches_prefix stringappend_8340)) of + Some (stringappend_8350,stringappend_8360) => (stringappend_8350, stringappend_8360) ) of - (_, stringappend_8000) => - (let stringappend_8010 = (string_drop stringappend_7980 stringappend_8000) in - (let (rs1, stringappend_8030) = - ((case ((reg_name_matches_prefix stringappend_8010 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8020,stringappend_8030) => (stringappend_8020, stringappend_8030) + (_, stringappend_8360) => + (let stringappend_8370 = (string_drop stringappend_8340 stringappend_8360) in + (let (rs1, stringappend_8390) = + ((case ((reg_name_matches_prefix stringappend_8370 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8380,stringappend_8390) => (stringappend_8380, stringappend_8390) )) in - (let stringappend_8040 = (string_drop stringappend_8010 stringappend_8030) in + (let stringappend_8400 = (string_drop stringappend_8370 stringappend_8390) in (case - (case ((sep_matches_prefix stringappend_8040)) of - Some (stringappend_8050,stringappend_8060) => (stringappend_8050, stringappend_8060) + (case ((sep_matches_prefix stringappend_8400)) of + Some (stringappend_8410,stringappend_8420) => (stringappend_8410, stringappend_8420) ) of - (_, stringappend_8060) => - (let stringappend_8070 = (string_drop stringappend_8040 stringappend_8060) in - (let (rs2, stringappend_8090) = - ((case ((reg_name_matches_prefix stringappend_8070 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8080,stringappend_8090) => (stringappend_8080, stringappend_8090) + (_, stringappend_8420) => + (let stringappend_8430 = (string_drop stringappend_8400 stringappend_8420) in + (let (rs2, stringappend_8450) = + ((case ((reg_name_matches_prefix stringappend_8430 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8440,stringappend_8450) => (stringappend_8440, stringappend_8450) )) in - if(((string_drop stringappend_8070 stringappend_8090)) = ('''')) then + if(((string_drop stringappend_8430 stringappend_8450)) = ('''')) then True else undefined)) )))) )))) ))) - else if (((((string_startswith stringappend_5410 (''div''))) \<and> ((let stringappend_8110 = (string_drop stringappend_5410 ((string_length (''div'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_8110)) of - Some (stringappend_8120,stringappend_8130) => - (let stringappend_8140 = (string_drop stringappend_8110 stringappend_8130) in - if ((case ((spc_matches_prefix stringappend_8140)) of - Some (stringappend_8150,stringappend_8160) => - (let stringappend_8170 = (string_drop stringappend_8140 stringappend_8160) in - if ((case ((reg_name_matches_prefix stringappend_8170 + else if (((((string_startswith stringappend_5590 (''div''))) \<and> ((let stringappend_8470 = (string_drop stringappend_5590 ((string_length (''div'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_8470)) of + Some (stringappend_8480,stringappend_8490) => + (let stringappend_8500 = (string_drop stringappend_8470 stringappend_8490) in + if ((case ((spc_matches_prefix stringappend_8500)) of + Some (stringappend_8510,stringappend_8520) => + (let stringappend_8530 = (string_drop stringappend_8500 stringappend_8520) in + if ((case ((reg_name_matches_prefix stringappend_8530 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8180,stringappend_8190) => - (let stringappend_8200 = - (string_drop stringappend_8170 stringappend_8190) in - if ((case ((sep_matches_prefix stringappend_8200)) of - Some (stringappend_8210,stringappend_8220) => - (let stringappend_8230 = - (string_drop stringappend_8200 stringappend_8220) in - if ((case ((reg_name_matches_prefix stringappend_8230 + Some (stringappend_8540,stringappend_8550) => + (let stringappend_8560 = + (string_drop stringappend_8530 stringappend_8550) in + if ((case ((sep_matches_prefix stringappend_8560)) of + Some (stringappend_8570,stringappend_8580) => + (let stringappend_8590 = + (string_drop stringappend_8560 stringappend_8580) in + if ((case ((reg_name_matches_prefix stringappend_8590 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8240,stringappend_8250) => - (let stringappend_8260 = - (string_drop stringappend_8230 stringappend_8250) in - if ((case ((sep_matches_prefix stringappend_8260)) of - Some (stringappend_8270,stringappend_8280) => - (let stringappend_8290 = - (string_drop stringappend_8260 stringappend_8280) in - if ((case ((reg_name_matches_prefix stringappend_8290 + Some (stringappend_8600,stringappend_8610) => + (let stringappend_8620 = + (string_drop stringappend_8590 stringappend_8610) in + if ((case ((sep_matches_prefix stringappend_8620)) of + Some (stringappend_8630,stringappend_8640) => + (let stringappend_8650 = + (string_drop stringappend_8620 stringappend_8640) in + if ((case ((reg_name_matches_prefix stringappend_8650 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8300,stringappend_8310) => - if(((string_drop stringappend_8290 stringappend_8310)) = ('''')) then + Some (stringappend_8660,stringappend_8670) => + if(((string_drop stringappend_8650 stringappend_8670)) = ('''')) then True else False | None => False )) then @@ -18346,78 +18596,78 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_8110 = (string_drop stringappend_5410 ((string_length (''div'')))) in - (let (s, stringappend_8130) = - ((case ((maybe_not_u_matches_prefix stringappend_8110)) of - Some (stringappend_8120,stringappend_8130) => (stringappend_8120, stringappend_8130) + (let stringappend_8470 = (string_drop stringappend_5590 ((string_length (''div'')))) in + (let (s, stringappend_8490) = + ((case ((maybe_not_u_matches_prefix stringappend_8470)) of + Some (stringappend_8480,stringappend_8490) => (stringappend_8480, stringappend_8490) )) in - (let stringappend_8140 = (string_drop stringappend_8110 stringappend_8130) in + (let stringappend_8500 = (string_drop stringappend_8470 stringappend_8490) in (case - (case ((spc_matches_prefix stringappend_8140)) of - Some (stringappend_8150,stringappend_8160) => (stringappend_8150, stringappend_8160) + (case ((spc_matches_prefix stringappend_8500)) of + Some (stringappend_8510,stringappend_8520) => (stringappend_8510, stringappend_8520) ) of - (_, stringappend_8160) => - (let stringappend_8170 = (string_drop stringappend_8140 stringappend_8160) in - (let (rd, stringappend_8190) = - ((case ((reg_name_matches_prefix stringappend_8170 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8180,stringappend_8190) => (stringappend_8180, stringappend_8190) + (_, stringappend_8520) => + (let stringappend_8530 = (string_drop stringappend_8500 stringappend_8520) in + (let (rd, stringappend_8550) = + ((case ((reg_name_matches_prefix stringappend_8530 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8540,stringappend_8550) => (stringappend_8540, stringappend_8550) )) in - (let stringappend_8200 = (string_drop stringappend_8170 stringappend_8190) in + (let stringappend_8560 = (string_drop stringappend_8530 stringappend_8550) in (case - (case ((sep_matches_prefix stringappend_8200)) of - Some (stringappend_8210,stringappend_8220) => (stringappend_8210, stringappend_8220) + (case ((sep_matches_prefix stringappend_8560)) of + Some (stringappend_8570,stringappend_8580) => (stringappend_8570, stringappend_8580) ) of - (_, stringappend_8220) => - (let stringappend_8230 = (string_drop stringappend_8200 stringappend_8220) in - (let (rs1, stringappend_8250) = - ((case ((reg_name_matches_prefix stringappend_8230 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8240,stringappend_8250) => (stringappend_8240, stringappend_8250) + (_, stringappend_8580) => + (let stringappend_8590 = (string_drop stringappend_8560 stringappend_8580) in + (let (rs1, stringappend_8610) = + ((case ((reg_name_matches_prefix stringappend_8590 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8600,stringappend_8610) => (stringappend_8600, stringappend_8610) )) in - (let stringappend_8260 = (string_drop stringappend_8230 stringappend_8250) in + (let stringappend_8620 = (string_drop stringappend_8590 stringappend_8610) in (case - (case ((sep_matches_prefix stringappend_8260)) of - Some (stringappend_8270,stringappend_8280) => (stringappend_8270, stringappend_8280) + (case ((sep_matches_prefix stringappend_8620)) of + Some (stringappend_8630,stringappend_8640) => (stringappend_8630, stringappend_8640) ) of - (_, stringappend_8280) => - (let stringappend_8290 = (string_drop stringappend_8260 stringappend_8280) in - (let (rs2, stringappend_8310) = - ((case ((reg_name_matches_prefix stringappend_8290 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8300,stringappend_8310) => (stringappend_8300, stringappend_8310) + (_, stringappend_8640) => + (let stringappend_8650 = (string_drop stringappend_8620 stringappend_8640) in + (let (rs2, stringappend_8670) = + ((case ((reg_name_matches_prefix stringappend_8650 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8660,stringappend_8670) => (stringappend_8660, stringappend_8670) )) in - if(((string_drop stringappend_8290 stringappend_8310)) = ('''')) then + if(((string_drop stringappend_8650 stringappend_8670)) = ('''')) then True else undefined)) )))) )))) )))) - else if (((((string_startswith stringappend_5410 (''rem''))) \<and> ((let stringappend_8330 = (string_drop stringappend_5410 ((string_length (''rem'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_8330)) of - Some (stringappend_8340,stringappend_8350) => - (let stringappend_8360 = (string_drop stringappend_8330 stringappend_8350) in - if ((case ((spc_matches_prefix stringappend_8360)) of - Some (stringappend_8370,stringappend_8380) => - (let stringappend_8390 = (string_drop stringappend_8360 stringappend_8380) in - if ((case ((reg_name_matches_prefix stringappend_8390 + else if (((((string_startswith stringappend_5590 (''rem''))) \<and> ((let stringappend_8690 = (string_drop stringappend_5590 ((string_length (''rem'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_8690)) of + Some (stringappend_8700,stringappend_8710) => + (let stringappend_8720 = (string_drop stringappend_8690 stringappend_8710) in + if ((case ((spc_matches_prefix stringappend_8720)) of + Some (stringappend_8730,stringappend_8740) => + (let stringappend_8750 = (string_drop stringappend_8720 stringappend_8740) in + if ((case ((reg_name_matches_prefix stringappend_8750 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8400,stringappend_8410) => - (let stringappend_8420 = - (string_drop stringappend_8390 stringappend_8410) in - if ((case ((sep_matches_prefix stringappend_8420)) of - Some (stringappend_8430,stringappend_8440) => - (let stringappend_8450 = - (string_drop stringappend_8420 stringappend_8440) in - if ((case ((reg_name_matches_prefix stringappend_8450 + Some (stringappend_8760,stringappend_8770) => + (let stringappend_8780 = + (string_drop stringappend_8750 stringappend_8770) in + if ((case ((sep_matches_prefix stringappend_8780)) of + Some (stringappend_8790,stringappend_8800) => + (let stringappend_8810 = + (string_drop stringappend_8780 stringappend_8800) in + if ((case ((reg_name_matches_prefix stringappend_8810 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8460,stringappend_8470) => - (let stringappend_8480 = - (string_drop stringappend_8450 stringappend_8470) in - if ((case ((sep_matches_prefix stringappend_8480)) of - Some (stringappend_8490,stringappend_8500) => - (let stringappend_8510 = - (string_drop stringappend_8480 stringappend_8500) in - if ((case ((reg_name_matches_prefix stringappend_8510 + Some (stringappend_8820,stringappend_8830) => + (let stringappend_8840 = + (string_drop stringappend_8810 stringappend_8830) in + if ((case ((sep_matches_prefix stringappend_8840)) of + Some (stringappend_8850,stringappend_8860) => + (let stringappend_8870 = + (string_drop stringappend_8840 stringappend_8860) in + if ((case ((reg_name_matches_prefix stringappend_8870 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8520,stringappend_8530) => - if(((string_drop stringappend_8510 stringappend_8530)) = ('''')) then + Some (stringappend_8880,stringappend_8890) => + if(((string_drop stringappend_8870 stringappend_8890)) = ('''')) then True else False | None => False )) then @@ -18447,74 +18697,74 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_8330 = (string_drop stringappend_5410 ((string_length (''rem'')))) in - (let (s, stringappend_8350) = - ((case ((maybe_not_u_matches_prefix stringappend_8330)) of - Some (stringappend_8340,stringappend_8350) => (stringappend_8340, stringappend_8350) + (let stringappend_8690 = (string_drop stringappend_5590 ((string_length (''rem'')))) in + (let (s, stringappend_8710) = + ((case ((maybe_not_u_matches_prefix stringappend_8690)) of + Some (stringappend_8700,stringappend_8710) => (stringappend_8700, stringappend_8710) )) in - (let stringappend_8360 = (string_drop stringappend_8330 stringappend_8350) in + (let stringappend_8720 = (string_drop stringappend_8690 stringappend_8710) in (case - (case ((spc_matches_prefix stringappend_8360)) of - Some (stringappend_8370,stringappend_8380) => (stringappend_8370, stringappend_8380) + (case ((spc_matches_prefix stringappend_8720)) of + Some (stringappend_8730,stringappend_8740) => (stringappend_8730, stringappend_8740) ) of - (_, stringappend_8380) => - (let stringappend_8390 = (string_drop stringappend_8360 stringappend_8380) in - (let (rd, stringappend_8410) = - ((case ((reg_name_matches_prefix stringappend_8390 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8400,stringappend_8410) => (stringappend_8400, stringappend_8410) + (_, stringappend_8740) => + (let stringappend_8750 = (string_drop stringappend_8720 stringappend_8740) in + (let (rd, stringappend_8770) = + ((case ((reg_name_matches_prefix stringappend_8750 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8760,stringappend_8770) => (stringappend_8760, stringappend_8770) )) in - (let stringappend_8420 = (string_drop stringappend_8390 stringappend_8410) in + (let stringappend_8780 = (string_drop stringappend_8750 stringappend_8770) in (case - (case ((sep_matches_prefix stringappend_8420)) of - Some (stringappend_8430,stringappend_8440) => (stringappend_8430, stringappend_8440) + (case ((sep_matches_prefix stringappend_8780)) of + Some (stringappend_8790,stringappend_8800) => (stringappend_8790, stringappend_8800) ) of - (_, stringappend_8440) => - (let stringappend_8450 = (string_drop stringappend_8420 stringappend_8440) in - (let (rs1, stringappend_8470) = - ((case ((reg_name_matches_prefix stringappend_8450 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8460,stringappend_8470) => (stringappend_8460, stringappend_8470) + (_, stringappend_8800) => + (let stringappend_8810 = (string_drop stringappend_8780 stringappend_8800) in + (let (rs1, stringappend_8830) = + ((case ((reg_name_matches_prefix stringappend_8810 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8820,stringappend_8830) => (stringappend_8820, stringappend_8830) )) in - (let stringappend_8480 = (string_drop stringappend_8450 stringappend_8470) in + (let stringappend_8840 = (string_drop stringappend_8810 stringappend_8830) in (case - (case ((sep_matches_prefix stringappend_8480)) of - Some (stringappend_8490,stringappend_8500) => (stringappend_8490, stringappend_8500) + (case ((sep_matches_prefix stringappend_8840)) of + Some (stringappend_8850,stringappend_8860) => (stringappend_8850, stringappend_8860) ) of - (_, stringappend_8500) => - (let stringappend_8510 = (string_drop stringappend_8480 stringappend_8500) in - (let (rs2, stringappend_8530) = - ((case ((reg_name_matches_prefix stringappend_8510 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8520,stringappend_8530) => (stringappend_8520, stringappend_8530) + (_, stringappend_8860) => + (let stringappend_8870 = (string_drop stringappend_8840 stringappend_8860) in + (let (rs2, stringappend_8890) = + ((case ((reg_name_matches_prefix stringappend_8870 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8880,stringappend_8890) => (stringappend_8880, stringappend_8890) )) in - if(((string_drop stringappend_8510 stringappend_8530)) = ('''')) then + if(((string_drop stringappend_8870 stringappend_8890)) = ('''')) then True else undefined)) )))) )))) )))) - else if (((((string_startswith stringappend_5410 (''mulw''))) \<and> ((let stringappend_8550 = (string_drop stringappend_5410 ((string_length (''mulw'')))) in - if ((case ((spc_matches_prefix stringappend_8550)) of - Some (stringappend_8560,stringappend_8570) => - (let stringappend_8580 = (string_drop stringappend_8550 stringappend_8570) in - if ((case ((reg_name_matches_prefix stringappend_8580 + else if (((((string_startswith stringappend_5590 (''mulw''))) \<and> ((let stringappend_8910 = (string_drop stringappend_5590 ((string_length (''mulw'')))) in + if ((case ((spc_matches_prefix stringappend_8910)) of + Some (stringappend_8920,stringappend_8930) => + (let stringappend_8940 = (string_drop stringappend_8910 stringappend_8930) in + if ((case ((reg_name_matches_prefix stringappend_8940 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8590,stringappend_8600) => - (let stringappend_8610 = (string_drop stringappend_8580 stringappend_8600) in - if ((case ((sep_matches_prefix stringappend_8610)) of - Some (stringappend_8620,stringappend_8630) => - (let stringappend_8640 = - (string_drop stringappend_8610 stringappend_8630) in - if ((case ((reg_name_matches_prefix stringappend_8640 + Some (stringappend_8950,stringappend_8960) => + (let stringappend_8970 = (string_drop stringappend_8940 stringappend_8960) in + if ((case ((sep_matches_prefix stringappend_8970)) of + Some (stringappend_8980,stringappend_8990) => + (let stringappend_9000 = + (string_drop stringappend_8970 stringappend_8990) in + if ((case ((reg_name_matches_prefix stringappend_9000 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8650,stringappend_8660) => - (let stringappend_8670 = - (string_drop stringappend_8640 stringappend_8660) in - if ((case ((sep_matches_prefix stringappend_8670)) of - Some (stringappend_8680,stringappend_8690) => - (let stringappend_8700 = - (string_drop stringappend_8670 stringappend_8690) in - if ((case ((reg_name_matches_prefix stringappend_8700 + Some (stringappend_9010,stringappend_9020) => + (let stringappend_9030 = + (string_drop stringappend_9000 stringappend_9020) in + if ((case ((sep_matches_prefix stringappend_9030)) of + Some (stringappend_9040,stringappend_9050) => + (let stringappend_9060 = + (string_drop stringappend_9030 stringappend_9050) in + if ((case ((reg_name_matches_prefix stringappend_9060 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8710,stringappend_8720) => - if(((string_drop stringappend_8700 stringappend_8720)) = ('''')) then + Some (stringappend_9070,stringappend_9080) => + if(((string_drop stringappend_9060 stringappend_9080)) = ('''')) then True else False | None => False )) then @@ -18540,78 +18790,78 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_8550 = (string_drop stringappend_5410 ((string_length (''mulw'')))) in + (let stringappend_8910 = (string_drop stringappend_5590 ((string_length (''mulw'')))) in (case - (case ((spc_matches_prefix stringappend_8550)) of - Some (stringappend_8560,stringappend_8570) => (stringappend_8560, stringappend_8570) + (case ((spc_matches_prefix stringappend_8910)) of + Some (stringappend_8920,stringappend_8930) => (stringappend_8920, stringappend_8930) ) of - (_, stringappend_8570) => - (let stringappend_8580 = (string_drop stringappend_8550 stringappend_8570) in - (let (rd, stringappend_8600) = - ((case ((reg_name_matches_prefix stringappend_8580 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8590,stringappend_8600) => (stringappend_8590, stringappend_8600) + (_, stringappend_8930) => + (let stringappend_8940 = (string_drop stringappend_8910 stringappend_8930) in + (let (rd, stringappend_8960) = + ((case ((reg_name_matches_prefix stringappend_8940 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_8950,stringappend_8960) => (stringappend_8950, stringappend_8960) )) in - (let stringappend_8610 = (string_drop stringappend_8580 stringappend_8600) in + (let stringappend_8970 = (string_drop stringappend_8940 stringappend_8960) in (case - (case ((sep_matches_prefix stringappend_8610)) of - Some (stringappend_8620,stringappend_8630) => (stringappend_8620, stringappend_8630) + (case ((sep_matches_prefix stringappend_8970)) of + Some (stringappend_8980,stringappend_8990) => (stringappend_8980, stringappend_8990) ) of - (_, stringappend_8630) => - (let stringappend_8640 = (string_drop stringappend_8610 stringappend_8630) in - (let (rs1, stringappend_8660) = - ((case ((reg_name_matches_prefix stringappend_8640 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8650,stringappend_8660) => (stringappend_8650, stringappend_8660) + (_, stringappend_8990) => + (let stringappend_9000 = (string_drop stringappend_8970 stringappend_8990) in + (let (rs1, stringappend_9020) = + ((case ((reg_name_matches_prefix stringappend_9000 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9010,stringappend_9020) => (stringappend_9010, stringappend_9020) )) in - (let stringappend_8670 = (string_drop stringappend_8640 stringappend_8660) in + (let stringappend_9030 = (string_drop stringappend_9000 stringappend_9020) in (case - (case ((sep_matches_prefix stringappend_8670)) of - Some (stringappend_8680,stringappend_8690) => (stringappend_8680, stringappend_8690) + (case ((sep_matches_prefix stringappend_9030)) of + Some (stringappend_9040,stringappend_9050) => (stringappend_9040, stringappend_9050) ) of - (_, stringappend_8690) => - (let stringappend_8700 = (string_drop stringappend_8670 stringappend_8690) in - (let (rs2, stringappend_8720) = - ((case ((reg_name_matches_prefix stringappend_8700 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8710,stringappend_8720) => (stringappend_8710, stringappend_8720) + (_, stringappend_9050) => + (let stringappend_9060 = (string_drop stringappend_9030 stringappend_9050) in + (let (rs2, stringappend_9080) = + ((case ((reg_name_matches_prefix stringappend_9060 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9070,stringappend_9080) => (stringappend_9070, stringappend_9080) )) in - if(((string_drop stringappend_8700 stringappend_8720)) = ('''')) then + if(((string_drop stringappend_9060 stringappend_9080)) = ('''')) then True else undefined)) )))) )))) )) - else if (((((string_startswith stringappend_5410 (''div''))) \<and> ((let stringappend_8740 = (string_drop stringappend_5410 ((string_length (''div'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_8740)) of - Some (stringappend_8750,stringappend_8760) => - (let stringappend_8770 = (string_drop stringappend_8740 stringappend_8760) in - if (((((string_startswith stringappend_8770 (''w''))) \<and> ((let stringappend_8780 = - (string_drop stringappend_8770 ((string_length (''w'')))) in - if ((case ((spc_matches_prefix stringappend_8780)) of - Some (stringappend_8790,stringappend_8800) => - (let stringappend_8810 = - (string_drop stringappend_8780 stringappend_8800) in - if ((case ((reg_name_matches_prefix stringappend_8810 + else if (((((string_startswith stringappend_5590 (''div''))) \<and> ((let stringappend_9100 = (string_drop stringappend_5590 ((string_length (''div'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_9100)) of + Some (stringappend_9110,stringappend_9120) => + (let stringappend_9130 = (string_drop stringappend_9100 stringappend_9120) in + if (((((string_startswith stringappend_9130 (''w''))) \<and> ((let stringappend_9140 = + (string_drop stringappend_9130 ((string_length (''w'')))) in + if ((case ((spc_matches_prefix stringappend_9140)) of + Some (stringappend_9150,stringappend_9160) => + (let stringappend_9170 = + (string_drop stringappend_9140 stringappend_9160) in + if ((case ((reg_name_matches_prefix stringappend_9170 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8820,stringappend_8830) => - (let stringappend_8840 = - (string_drop stringappend_8810 stringappend_8830) in - if ((case ((sep_matches_prefix stringappend_8840)) of - Some (stringappend_8850,stringappend_8860) => - (let stringappend_8870 = - (string_drop stringappend_8840 stringappend_8860) in - if ((case ((reg_name_matches_prefix stringappend_8870 + Some (stringappend_9180,stringappend_9190) => + (let stringappend_9200 = + (string_drop stringappend_9170 stringappend_9190) in + if ((case ((sep_matches_prefix stringappend_9200)) of + Some (stringappend_9210,stringappend_9220) => + (let stringappend_9230 = + (string_drop stringappend_9200 stringappend_9220) in + if ((case ((reg_name_matches_prefix stringappend_9230 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8880,stringappend_8890) => - (let stringappend_8900 = - (string_drop stringappend_8870 stringappend_8890) in - if ((case ((sep_matches_prefix stringappend_8900)) of - Some (stringappend_8910,stringappend_8920) => - (let stringappend_8930 = - (string_drop stringappend_8900 - stringappend_8920) in + Some (stringappend_9240,stringappend_9250) => + (let stringappend_9260 = + (string_drop stringappend_9230 stringappend_9250) in + if ((case ((sep_matches_prefix stringappend_9260)) of + Some (stringappend_9270,stringappend_9280) => + (let stringappend_9290 = + (string_drop stringappend_9260 + stringappend_9280) in if ((case ((reg_name_matches_prefix - stringappend_8930 + stringappend_9290 :: (( 5 Word.word * ii))option)) of - Some (stringappend_8940,stringappend_8950) => - if(((string_drop stringappend_8930 stringappend_8950)) = ('''')) then + Some (stringappend_9300,stringappend_9310) => + if(((string_drop stringappend_9290 stringappend_9310)) = ('''')) then True else False | None => False )) then @@ -18643,84 +18893,84 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_8740 = (string_drop stringappend_5410 ((string_length (''div'')))) in - (let (s, stringappend_8760) = - ((case ((maybe_not_u_matches_prefix stringappend_8740)) of - Some (stringappend_8750,stringappend_8760) => (stringappend_8750, stringappend_8760) + (let stringappend_9100 = (string_drop stringappend_5590 ((string_length (''div'')))) in + (let (s, stringappend_9120) = + ((case ((maybe_not_u_matches_prefix stringappend_9100)) of + Some (stringappend_9110,stringappend_9120) => (stringappend_9110, stringappend_9120) )) in - (let stringappend_8770 = (string_drop stringappend_8740 stringappend_8760) in - (let stringappend_8780 = (string_drop stringappend_8770 ((string_length (''w'')))) in + (let stringappend_9130 = (string_drop stringappend_9100 stringappend_9120) in + (let stringappend_9140 = (string_drop stringappend_9130 ((string_length (''w'')))) in (case - (case ((spc_matches_prefix stringappend_8780)) of - Some (stringappend_8790,stringappend_8800) => (stringappend_8790, stringappend_8800) + (case ((spc_matches_prefix stringappend_9140)) of + Some (stringappend_9150,stringappend_9160) => (stringappend_9150, stringappend_9160) ) of - (_, stringappend_8800) => - (let stringappend_8810 = (string_drop stringappend_8780 stringappend_8800) in - (let (rd, stringappend_8830) = - ((case ((reg_name_matches_prefix stringappend_8810 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8820,stringappend_8830) => (stringappend_8820, stringappend_8830) + (_, stringappend_9160) => + (let stringappend_9170 = (string_drop stringappend_9140 stringappend_9160) in + (let (rd, stringappend_9190) = + ((case ((reg_name_matches_prefix stringappend_9170 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9180,stringappend_9190) => (stringappend_9180, stringappend_9190) )) in - (let stringappend_8840 = (string_drop stringappend_8810 stringappend_8830) in + (let stringappend_9200 = (string_drop stringappend_9170 stringappend_9190) in (case - (case ((sep_matches_prefix stringappend_8840)) of - Some (stringappend_8850,stringappend_8860) => (stringappend_8850, stringappend_8860) + (case ((sep_matches_prefix stringappend_9200)) of + Some (stringappend_9210,stringappend_9220) => (stringappend_9210, stringappend_9220) ) of - (_, stringappend_8860) => - (let stringappend_8870 = (string_drop stringappend_8840 stringappend_8860) in - (let (rs1, stringappend_8890) = - ((case ((reg_name_matches_prefix stringappend_8870 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8880,stringappend_8890) => (stringappend_8880, stringappend_8890) + (_, stringappend_9220) => + (let stringappend_9230 = (string_drop stringappend_9200 stringappend_9220) in + (let (rs1, stringappend_9250) = + ((case ((reg_name_matches_prefix stringappend_9230 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9240,stringappend_9250) => (stringappend_9240, stringappend_9250) )) in - (let stringappend_8900 = (string_drop stringappend_8870 stringappend_8890) in + (let stringappend_9260 = (string_drop stringappend_9230 stringappend_9250) in (case - (case ((sep_matches_prefix stringappend_8900)) of - Some (stringappend_8910,stringappend_8920) => (stringappend_8910, stringappend_8920) + (case ((sep_matches_prefix stringappend_9260)) of + Some (stringappend_9270,stringappend_9280) => (stringappend_9270, stringappend_9280) ) of - (_, stringappend_8920) => - (let stringappend_8930 = (string_drop stringappend_8900 stringappend_8920) in - (let (rs2, stringappend_8950) = - ((case ((reg_name_matches_prefix stringappend_8930 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_8940,stringappend_8950) => (stringappend_8940, stringappend_8950) + (_, stringappend_9280) => + (let stringappend_9290 = (string_drop stringappend_9260 stringappend_9280) in + (let (rs2, stringappend_9310) = + ((case ((reg_name_matches_prefix stringappend_9290 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9300,stringappend_9310) => (stringappend_9300, stringappend_9310) )) in - if(((string_drop stringappend_8930 stringappend_8950)) = ('''')) then + if(((string_drop stringappend_9290 stringappend_9310)) = ('''')) then True else undefined)) )))) )))) ))))) - else if (((((string_startswith stringappend_5410 (''rem''))) \<and> ((let stringappend_8970 = (string_drop stringappend_5410 ((string_length (''rem'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_8970)) of - Some (stringappend_8980,stringappend_8990) => - (let stringappend_9000 = (string_drop stringappend_8970 stringappend_8990) in - if (((((string_startswith stringappend_9000 (''w''))) \<and> ((let stringappend_9010 = - (string_drop stringappend_9000 ((string_length (''w'')))) in - if ((case ((spc_matches_prefix stringappend_9010)) of - Some (stringappend_9020,stringappend_9030) => - (let stringappend_9040 = - (string_drop stringappend_9010 stringappend_9030) in - if ((case ((reg_name_matches_prefix stringappend_9040 + else if (((((string_startswith stringappend_5590 (''rem''))) \<and> ((let stringappend_9330 = (string_drop stringappend_5590 ((string_length (''rem'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_9330)) of + Some (stringappend_9340,stringappend_9350) => + (let stringappend_9360 = (string_drop stringappend_9330 stringappend_9350) in + if (((((string_startswith stringappend_9360 (''w''))) \<and> ((let stringappend_9370 = + (string_drop stringappend_9360 ((string_length (''w'')))) in + if ((case ((spc_matches_prefix stringappend_9370)) of + Some (stringappend_9380,stringappend_9390) => + (let stringappend_9400 = + (string_drop stringappend_9370 stringappend_9390) in + if ((case ((reg_name_matches_prefix stringappend_9400 :: (( 5 Word.word * ii))option)) of - Some (stringappend_9050,stringappend_9060) => - (let stringappend_9070 = - (string_drop stringappend_9040 stringappend_9060) in - if ((case ((sep_matches_prefix stringappend_9070)) of - Some (stringappend_9080,stringappend_9090) => - (let stringappend_9100 = - (string_drop stringappend_9070 stringappend_9090) in - if ((case ((reg_name_matches_prefix stringappend_9100 + Some (stringappend_9410,stringappend_9420) => + (let stringappend_9430 = + (string_drop stringappend_9400 stringappend_9420) in + if ((case ((sep_matches_prefix stringappend_9430)) of + Some (stringappend_9440,stringappend_9450) => + (let stringappend_9460 = + (string_drop stringappend_9430 stringappend_9450) in + if ((case ((reg_name_matches_prefix stringappend_9460 :: (( 5 Word.word * ii))option)) of - Some (stringappend_9110,stringappend_9120) => - (let stringappend_9130 = - (string_drop stringappend_9100 stringappend_9120) in - if ((case ((sep_matches_prefix stringappend_9130)) of - Some (stringappend_9140,stringappend_9150) => - (let stringappend_9160 = - (string_drop stringappend_9130 - stringappend_9150) in + Some (stringappend_9470,stringappend_9480) => + (let stringappend_9490 = + (string_drop stringappend_9460 stringappend_9480) in + if ((case ((sep_matches_prefix stringappend_9490)) of + Some (stringappend_9500,stringappend_9510) => + (let stringappend_9520 = + (string_drop stringappend_9490 + stringappend_9510) in if ((case ((reg_name_matches_prefix - stringappend_9160 + stringappend_9520 :: (( 5 Word.word * ii))option)) of - Some (stringappend_9170,stringappend_9180) => - if(((string_drop stringappend_9160 stringappend_9180)) = ('''')) then + Some (stringappend_9530,stringappend_9540) => + if(((string_drop stringappend_9520 stringappend_9540)) = ('''')) then True else False | None => False )) then @@ -18752,66 +19002,66 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_8970 = (string_drop stringappend_5410 ((string_length (''rem'')))) in - (let (s, stringappend_8990) = - ((case ((maybe_not_u_matches_prefix stringappend_8970)) of - Some (stringappend_8980,stringappend_8990) => (stringappend_8980, stringappend_8990) + (let stringappend_9330 = (string_drop stringappend_5590 ((string_length (''rem'')))) in + (let (s, stringappend_9350) = + ((case ((maybe_not_u_matches_prefix stringappend_9330)) of + Some (stringappend_9340,stringappend_9350) => (stringappend_9340, stringappend_9350) )) in - (let stringappend_9000 = (string_drop stringappend_8970 stringappend_8990) in - (let stringappend_9010 = (string_drop stringappend_9000 ((string_length (''w'')))) in + (let stringappend_9360 = (string_drop stringappend_9330 stringappend_9350) in + (let stringappend_9370 = (string_drop stringappend_9360 ((string_length (''w'')))) in (case - (case ((spc_matches_prefix stringappend_9010)) of - Some (stringappend_9020,stringappend_9030) => (stringappend_9020, stringappend_9030) + (case ((spc_matches_prefix stringappend_9370)) of + Some (stringappend_9380,stringappend_9390) => (stringappend_9380, stringappend_9390) ) of - (_, stringappend_9030) => - (let stringappend_9040 = (string_drop stringappend_9010 stringappend_9030) in - (let (rd, stringappend_9060) = - ((case ((reg_name_matches_prefix stringappend_9040 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9050,stringappend_9060) => (stringappend_9050, stringappend_9060) + (_, stringappend_9390) => + (let stringappend_9400 = (string_drop stringappend_9370 stringappend_9390) in + (let (rd, stringappend_9420) = + ((case ((reg_name_matches_prefix stringappend_9400 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9410,stringappend_9420) => (stringappend_9410, stringappend_9420) )) in - (let stringappend_9070 = (string_drop stringappend_9040 stringappend_9060) in + (let stringappend_9430 = (string_drop stringappend_9400 stringappend_9420) in (case - (case ((sep_matches_prefix stringappend_9070)) of - Some (stringappend_9080,stringappend_9090) => (stringappend_9080, stringappend_9090) + (case ((sep_matches_prefix stringappend_9430)) of + Some (stringappend_9440,stringappend_9450) => (stringappend_9440, stringappend_9450) ) of - (_, stringappend_9090) => - (let stringappend_9100 = (string_drop stringappend_9070 stringappend_9090) in - (let (rs1, stringappend_9120) = - ((case ((reg_name_matches_prefix stringappend_9100 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9110,stringappend_9120) => (stringappend_9110, stringappend_9120) + (_, stringappend_9450) => + (let stringappend_9460 = (string_drop stringappend_9430 stringappend_9450) in + (let (rs1, stringappend_9480) = + ((case ((reg_name_matches_prefix stringappend_9460 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9470,stringappend_9480) => (stringappend_9470, stringappend_9480) )) in - (let stringappend_9130 = (string_drop stringappend_9100 stringappend_9120) in + (let stringappend_9490 = (string_drop stringappend_9460 stringappend_9480) in (case - (case ((sep_matches_prefix stringappend_9130)) of - Some (stringappend_9140,stringappend_9150) => (stringappend_9140, stringappend_9150) + (case ((sep_matches_prefix stringappend_9490)) of + Some (stringappend_9500,stringappend_9510) => (stringappend_9500, stringappend_9510) ) of - (_, stringappend_9150) => - (let stringappend_9160 = (string_drop stringappend_9130 stringappend_9150) in - (let (rs2, stringappend_9180) = - ((case ((reg_name_matches_prefix stringappend_9160 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9170,stringappend_9180) => (stringappend_9170, stringappend_9180) + (_, stringappend_9510) => + (let stringappend_9520 = (string_drop stringappend_9490 stringappend_9510) in + (let (rs2, stringappend_9540) = + ((case ((reg_name_matches_prefix stringappend_9520 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9530,stringappend_9540) => (stringappend_9530, stringappend_9540) )) in - if(((string_drop stringappend_9160 stringappend_9180)) = ('''')) then + if(((string_drop stringappend_9520 stringappend_9540)) = ('''')) then True else undefined)) )))) )))) ))))) - else if (((((string_startswith stringappend_5410 (''fence''))) \<and> ((let stringappend_9200 = (string_drop stringappend_5410 ((string_length (''fence'')))) in - if ((case ((spc_matches_prefix stringappend_9200)) of - Some (stringappend_9210,stringappend_9220) => - (let stringappend_9230 = (string_drop stringappend_9200 stringappend_9220) in - if ((case ((fence_bits_matches_prefix stringappend_9230 + else if (((((string_startswith stringappend_5590 (''fence''))) \<and> ((let stringappend_9560 = (string_drop stringappend_5590 ((string_length (''fence'')))) in + if ((case ((spc_matches_prefix stringappend_9560)) of + Some (stringappend_9570,stringappend_9580) => + (let stringappend_9590 = (string_drop stringappend_9560 stringappend_9580) in + if ((case ((fence_bits_matches_prefix stringappend_9590 :: (( 4 Word.word * ii))option)) of - Some (stringappend_9240,stringappend_9250) => - (let stringappend_9260 = (string_drop stringappend_9230 stringappend_9250) in - if ((case ((sep_matches_prefix stringappend_9260)) of - Some (stringappend_9270,stringappend_9280) => - (let stringappend_9290 = - (string_drop stringappend_9260 stringappend_9280) in - if ((case ((fence_bits_matches_prefix stringappend_9290 + Some (stringappend_9600,stringappend_9610) => + (let stringappend_9620 = (string_drop stringappend_9590 stringappend_9610) in + if ((case ((sep_matches_prefix stringappend_9620)) of + Some (stringappend_9630,stringappend_9640) => + (let stringappend_9650 = + (string_drop stringappend_9620 stringappend_9640) in + if ((case ((fence_bits_matches_prefix stringappend_9650 :: (( 4 Word.word * ii))option)) of - Some (stringappend_9300,stringappend_9310) => - if(((string_drop stringappend_9290 stringappend_9310)) = ('''')) then + Some (stringappend_9660,stringappend_9670) => + if(((string_drop stringappend_9650 stringappend_9670)) = ('''')) then True else False | None => False )) then @@ -18829,73 +19079,73 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False))))) then - (let stringappend_9200 = (string_drop stringappend_5410 ((string_length (''fence'')))) in + (let stringappend_9560 = (string_drop stringappend_5590 ((string_length (''fence'')))) in (case - (case ((spc_matches_prefix stringappend_9200)) of - Some (stringappend_9210,stringappend_9220) => (stringappend_9210, stringappend_9220) + (case ((spc_matches_prefix stringappend_9560)) of + Some (stringappend_9570,stringappend_9580) => (stringappend_9570, stringappend_9580) ) of - (_, stringappend_9220) => - (let stringappend_9230 = (string_drop stringappend_9200 stringappend_9220) in - (let (pred, stringappend_9250) = - ((case ((fence_bits_matches_prefix stringappend_9230 :: (( 4 Word.word * ii)) option)) of - Some (stringappend_9240,stringappend_9250) => (stringappend_9240, stringappend_9250) + (_, stringappend_9580) => + (let stringappend_9590 = (string_drop stringappend_9560 stringappend_9580) in + (let (pred, stringappend_9610) = + ((case ((fence_bits_matches_prefix stringappend_9590 :: (( 4 Word.word * ii)) option)) of + Some (stringappend_9600,stringappend_9610) => (stringappend_9600, stringappend_9610) )) in - (let stringappend_9260 = (string_drop stringappend_9230 stringappend_9250) in + (let stringappend_9620 = (string_drop stringappend_9590 stringappend_9610) in (case - (case ((sep_matches_prefix stringappend_9260)) of - Some (stringappend_9270,stringappend_9280) => (stringappend_9270, stringappend_9280) + (case ((sep_matches_prefix stringappend_9620)) of + Some (stringappend_9630,stringappend_9640) => (stringappend_9630, stringappend_9640) ) of - (_, stringappend_9280) => - (let stringappend_9290 = (string_drop stringappend_9260 stringappend_9280) in - (let (succ, stringappend_9310) = - ((case ((fence_bits_matches_prefix stringappend_9290 :: (( 4 Word.word * ii)) option)) of - Some (stringappend_9300,stringappend_9310) => (stringappend_9300, stringappend_9310) + (_, stringappend_9640) => + (let stringappend_9650 = (string_drop stringappend_9620 stringappend_9640) in + (let (succ, stringappend_9670) = + ((case ((fence_bits_matches_prefix stringappend_9650 :: (( 4 Word.word * ii)) option)) of + Some (stringappend_9660,stringappend_9670) => (stringappend_9660, stringappend_9670) )) in - if(((string_drop stringappend_9290 stringappend_9310)) = ('''')) then + if(((string_drop stringappend_9650 stringappend_9670)) = ('''')) then True else undefined)) )))) )) - else if(stringappend_5410 = (''fence.i'')) then True else + else if(stringappend_5590 = (''fence.i'')) then True else ( - if(stringappend_5410 = (''ecall'')) then True else + if(stringappend_5590 = (''ecall'')) then True else ( - if(stringappend_5410 = (''mret'')) then True else + if(stringappend_5590 = (''mret'')) then True else ( - if(stringappend_5410 = (''sret'')) then True else + if(stringappend_5590 = (''sret'')) then True else ( - if(stringappend_5410 = (''ebreak'')) then True else + if(stringappend_5590 = (''ebreak'')) then True else ( - if(stringappend_5410 = (''wfi'')) then True else + if(stringappend_5590 = (''wfi'')) then True else ( - if (((((string_startswith stringappend_5410 (''sfence.vma''))) + if (((((string_startswith stringappend_5590 (''sfence.vma''))) \<and> - ((let stringappend_9330 = - (string_drop stringappend_5410 + ((let stringappend_9690 = + (string_drop stringappend_5590 ((string_length (''sfence.vma'')))) in - if ((case ((spc_matches_prefix stringappend_9330)) of - Some (stringappend_9340,stringappend_9350) => - (let stringappend_9360 = (string_drop - stringappend_9330 - stringappend_9350) in + if ((case ((spc_matches_prefix stringappend_9690)) of + Some (stringappend_9700,stringappend_9710) => + (let stringappend_9720 = (string_drop + stringappend_9690 + stringappend_9710) in if ((case ((reg_name_matches_prefix - stringappend_9360 + stringappend_9720 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9370,stringappend_9380) => - (let stringappend_9390 = (string_drop - stringappend_9360 - stringappend_9380) in + Some (stringappend_9730,stringappend_9740) => + (let stringappend_9750 = (string_drop + stringappend_9720 + stringappend_9740) in if ((case ((sep_matches_prefix - stringappend_9390)) of - Some (stringappend_9400,stringappend_9410) => - (let stringappend_9420 = - (string_drop stringappend_9390 - stringappend_9410) in + stringappend_9750)) of + Some (stringappend_9760,stringappend_9770) => + (let stringappend_9780 = + (string_drop stringappend_9750 + stringappend_9770) in if ((case ((reg_name_matches_prefix - stringappend_9420 + stringappend_9780 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9430,stringappend_9440) => - if(((string_drop stringappend_9420 - stringappend_9440)) = ('''')) then + Some (stringappend_9790,stringappend_9800) => + if(((string_drop stringappend_9780 + stringappend_9800)) = ('''')) then True else False | None => False )) then True else False) @@ -18905,90 +19155,90 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False) | None => False )) then True else False))))) then - (let stringappend_9330 = (string_drop stringappend_5410 + (let stringappend_9690 = (string_drop stringappend_5590 ((string_length (''sfence.vma'')))) in (case - (case ((spc_matches_prefix stringappend_9330)) of - Some (stringappend_9340,stringappend_9350) => - (stringappend_9340, stringappend_9350) + (case ((spc_matches_prefix stringappend_9690)) of + Some (stringappend_9700,stringappend_9710) => + (stringappend_9700, stringappend_9710) ) of - (_, stringappend_9350) => - (let stringappend_9360 = (string_drop stringappend_9330 - stringappend_9350) in - (let (rs1, stringappend_9380) = - ((case ((reg_name_matches_prefix stringappend_9360 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9370,stringappend_9380) => - (stringappend_9370, stringappend_9380) + (_, stringappend_9710) => + (let stringappend_9720 = (string_drop stringappend_9690 + stringappend_9710) in + (let (rs1, stringappend_9740) = + ((case ((reg_name_matches_prefix stringappend_9720 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9730,stringappend_9740) => + (stringappend_9730, stringappend_9740) )) in - (let stringappend_9390 = (string_drop stringappend_9360 - stringappend_9380) in + (let stringappend_9750 = (string_drop stringappend_9720 + stringappend_9740) in (case - (case ((sep_matches_prefix stringappend_9390)) of - Some (stringappend_9400,stringappend_9410) => - (stringappend_9400, stringappend_9410) + (case ((sep_matches_prefix stringappend_9750)) of + Some (stringappend_9760,stringappend_9770) => + (stringappend_9760, stringappend_9770) ) of - (_, stringappend_9410) => - (let stringappend_9420 = (string_drop stringappend_9390 - stringappend_9410) in - (let (rs2, stringappend_9440) = - ((case ((reg_name_matches_prefix stringappend_9420 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9430,stringappend_9440) => - (stringappend_9430, stringappend_9440) + (_, stringappend_9770) => + (let stringappend_9780 = (string_drop stringappend_9750 + stringappend_9770) in + (let (rs2, stringappend_9800) = + ((case ((reg_name_matches_prefix stringappend_9780 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9790,stringappend_9800) => + (stringappend_9790, stringappend_9800) )) in - if(((string_drop stringappend_9420 stringappend_9440)) = + if(((string_drop stringappend_9780 stringappend_9800)) = ('''')) then True else undefined)) )))) )) else - if (((((string_startswith stringappend_5410 (''lr.''))) \<and> - ((let stringappend_9460 = (string_drop - stringappend_5410 + if (((((string_startswith stringappend_5590 (''lr.''))) \<and> + ((let stringappend_9820 = (string_drop + stringappend_5590 ((string_length (''lr.'')))) in if ((case ((maybe_aq_matches_prefix - stringappend_9460)) of - Some (stringappend_9470,stringappend_9480) => - (let stringappend_9490 = (string_drop - stringappend_9460 - stringappend_9480) in + stringappend_9820)) of + Some (stringappend_9830,stringappend_9840) => + (let stringappend_9850 = (string_drop + stringappend_9820 + stringappend_9840) in if ((case ((maybe_rl_matches_prefix - stringappend_9490)) of - Some (stringappend_9500,stringappend_9510) => - (let stringappend_9520 = - (string_drop stringappend_9490 - stringappend_9510) in + stringappend_9850)) of + Some (stringappend_9860,stringappend_9870) => + (let stringappend_9880 = + (string_drop stringappend_9850 + stringappend_9870) in if ((case ((size_mnemonic_matches_prefix - stringappend_9520)) of - Some (stringappend_9530,stringappend_9540) => - (let stringappend_9550 = - (string_drop stringappend_9520 - stringappend_9540) in + stringappend_9880)) of + Some (stringappend_9890,stringappend_9900) => + (let stringappend_9910 = + (string_drop stringappend_9880 + stringappend_9900) in if ((case ((spc_matches_prefix - stringappend_9550)) of - Some (stringappend_9560,stringappend_9570) => - (let stringappend_9580 = - (string_drop stringappend_9550 - stringappend_9570) in + stringappend_9910)) of + Some (stringappend_9920,stringappend_9930) => + (let stringappend_9940 = + (string_drop stringappend_9910 + stringappend_9930) in if ((case ((reg_name_matches_prefix - stringappend_9580 + stringappend_9940 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9590,stringappend_9600) => - (let stringappend_9610 = + Some (stringappend_9950,stringappend_9960) => + (let stringappend_9970 = (string_drop - stringappend_9580 - stringappend_9600) in + stringappend_9940 + stringappend_9960) in if ((case ((sep_matches_prefix - stringappend_9610)) of - Some (stringappend_9620,stringappend_9630) => - (let stringappend_9640 = + stringappend_9970)) of + Some (stringappend_9980,stringappend_9990) => + (let stringappend_10000 = (string_drop - stringappend_9610 - stringappend_9630) in + stringappend_9970 + stringappend_9990) in if ((case ((reg_name_matches_prefix - stringappend_9640 + stringappend_10000 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9650,stringappend_9660) => + Some (stringappend_10010,stringappend_10020) => if(((string_drop - stringappend_9640 - stringappend_9660)) + stringappend_10000 + stringappend_10020)) = ('''')) then True else False | None => False @@ -19006,132 +19256,132 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False) | None => False )) then True else False))))) then - (let stringappend_9460 = (string_drop stringappend_5410 + (let stringappend_9820 = (string_drop stringappend_5590 ((string_length (''lr.'')))) in - (let (aq, stringappend_9480) = - ((case ((maybe_aq_matches_prefix stringappend_9460)) of - Some (stringappend_9470,stringappend_9480) => - (stringappend_9470, stringappend_9480) + (let (aq, stringappend_9840) = + ((case ((maybe_aq_matches_prefix stringappend_9820)) of + Some (stringappend_9830,stringappend_9840) => + (stringappend_9830, stringappend_9840) )) in - (let stringappend_9490 = (string_drop stringappend_9460 - stringappend_9480) in - (let (rl, stringappend_9510) = - ((case ((maybe_rl_matches_prefix stringappend_9490)) of - Some (stringappend_9500,stringappend_9510) => - (stringappend_9500, stringappend_9510) + (let stringappend_9850 = (string_drop stringappend_9820 + stringappend_9840) in + (let (rl, stringappend_9870) = + ((case ((maybe_rl_matches_prefix stringappend_9850)) of + Some (stringappend_9860,stringappend_9870) => + (stringappend_9860, stringappend_9870) )) in - (let stringappend_9520 = (string_drop stringappend_9490 - stringappend_9510) in - (let (size1, stringappend_9540) = + (let stringappend_9880 = (string_drop stringappend_9850 + stringappend_9870) in + (let (size1, stringappend_9900) = ((case ((size_mnemonic_matches_prefix - stringappend_9520)) of - Some (stringappend_9530,stringappend_9540) => - (stringappend_9530, stringappend_9540) + stringappend_9880)) of + Some (stringappend_9890,stringappend_9900) => + (stringappend_9890, stringappend_9900) )) in - (let stringappend_9550 = (string_drop stringappend_9520 - stringappend_9540) in + (let stringappend_9910 = (string_drop stringappend_9880 + stringappend_9900) in (case - (case ((spc_matches_prefix stringappend_9550)) of - Some (stringappend_9560,stringappend_9570) => - (stringappend_9560, stringappend_9570) + (case ((spc_matches_prefix stringappend_9910)) of + Some (stringappend_9920,stringappend_9930) => + (stringappend_9920, stringappend_9930) ) of - (_, stringappend_9570) => - (let stringappend_9580 = (string_drop stringappend_9550 - stringappend_9570) in - (let (rd, stringappend_9600) = - ((case ((reg_name_matches_prefix stringappend_9580 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9590,stringappend_9600) => - (stringappend_9590, stringappend_9600) + (_, stringappend_9930) => + (let stringappend_9940 = (string_drop stringappend_9910 + stringappend_9930) in + (let (rd, stringappend_9960) = + ((case ((reg_name_matches_prefix stringappend_9940 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_9950,stringappend_9960) => + (stringappend_9950, stringappend_9960) )) in - (let stringappend_9610 = (string_drop stringappend_9580 - stringappend_9600) in + (let stringappend_9970 = (string_drop stringappend_9940 + stringappend_9960) in (case - (case ((sep_matches_prefix stringappend_9610)) of - Some (stringappend_9620,stringappend_9630) => - (stringappend_9620, stringappend_9630) + (case ((sep_matches_prefix stringappend_9970)) of + Some (stringappend_9980,stringappend_9990) => + (stringappend_9980, stringappend_9990) ) of - (_, stringappend_9630) => - (let stringappend_9640 = (string_drop stringappend_9610 - stringappend_9630) in - (let (rs1, stringappend_9660) = - ((case ((reg_name_matches_prefix stringappend_9640 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9650,stringappend_9660) => - (stringappend_9650, stringappend_9660) + (_, stringappend_9990) => + (let stringappend_10000 = (string_drop stringappend_9970 + stringappend_9990) in + (let (rs1, stringappend_10020) = + ((case ((reg_name_matches_prefix stringappend_10000 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10010,stringappend_10020) => + (stringappend_10010, stringappend_10020) )) in - if(((string_drop stringappend_9640 stringappend_9660)) = + if(((string_drop stringappend_10000 stringappend_10020)) = ('''')) then True else undefined)) )))) )))))))) else - if (((((string_startswith stringappend_5410 (''sc.''))) + if (((((string_startswith stringappend_5590 (''sc.''))) \<and> - ((let stringappend_9680 = (string_drop - stringappend_5410 - ((string_length - (''sc.'')))) in + ((let stringappend_10040 = + (string_drop stringappend_5590 + ((string_length (''sc.'')))) in if ((case ((maybe_aq_matches_prefix - stringappend_9680)) of - Some (stringappend_9690,stringappend_9700) => - (let stringappend_9710 = (string_drop - stringappend_9680 - stringappend_9700) in + stringappend_10040)) of + Some (stringappend_10050,stringappend_10060) => + (let stringappend_10070 = + (string_drop stringappend_10040 + stringappend_10060) in if ((case ((maybe_rl_matches_prefix - stringappend_9710)) of - Some (stringappend_9720,stringappend_9730) => - (let stringappend_9740 = - (string_drop stringappend_9710 - stringappend_9730) in + stringappend_10070)) of + Some (stringappend_10080,stringappend_10090) => + (let stringappend_10100 = + (string_drop stringappend_10070 + stringappend_10090) in if ((case ((size_mnemonic_matches_prefix - stringappend_9740)) of - Some (stringappend_9750,stringappend_9760) => - (let stringappend_9770 = - (string_drop stringappend_9740 - stringappend_9760) in + stringappend_10100)) of + Some (stringappend_10110,stringappend_10120) => + (let stringappend_10130 = + (string_drop stringappend_10100 + stringappend_10120) in if ((case ((spc_matches_prefix - stringappend_9770)) of - Some (stringappend_9780,stringappend_9790) => - (let stringappend_9800 = - (string_drop stringappend_9770 - stringappend_9790) in + stringappend_10130)) of + Some (stringappend_10140,stringappend_10150) => + (let stringappend_10160 = + (string_drop + stringappend_10130 + stringappend_10150) in if ((case ((reg_name_matches_prefix - stringappend_9800 + stringappend_10160 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9810,stringappend_9820) => - (let stringappend_9830 = + Some (stringappend_10170,stringappend_10180) => + (let stringappend_10190 = (string_drop - stringappend_9800 - stringappend_9820) in + stringappend_10160 + stringappend_10180) in if ((case ((sep_matches_prefix - stringappend_9830)) of - Some (stringappend_9840,stringappend_9850) => - (let stringappend_9860 = + stringappend_10190)) of + Some (stringappend_10200,stringappend_10210) => + (let stringappend_10220 = (string_drop - stringappend_9830 - stringappend_9850) in + stringappend_10190 + stringappend_10210) in if ((case ((reg_name_matches_prefix - stringappend_9860 + stringappend_10220 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9870,stringappend_9880) => - (let stringappend_9890 = + Some (stringappend_10230,stringappend_10240) => + (let stringappend_10250 = (string_drop - stringappend_9860 - stringappend_9880) in + stringappend_10220 + stringappend_10240) in if ((case ((sep_matches_prefix - stringappend_9890)) of + stringappend_10250)) of Some - (stringappend_9900,stringappend_9910) => - (let stringappend_9920 = + (stringappend_10260,stringappend_10270) => + (let stringappend_10280 = (string_drop - stringappend_9890 - stringappend_9910) in + stringappend_10250 + stringappend_10270) in if ((case ((reg_name_matches_prefix - stringappend_9920 + stringappend_10280 :: (( 5 Word.word * ii)) option)) of Some - (stringappend_9930,stringappend_9940) => + (stringappend_10290,stringappend_10300) => if(((string_drop - stringappend_9920 - stringappend_9940)) + stringappend_10280 + stringappend_10300)) = ('''')) then True else @@ -19159,164 +19409,164 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False) | None => False )) then True else False))))) then - (let stringappend_9680 = (string_drop stringappend_5410 - ((string_length (''sc.'')))) in - (let (aq, stringappend_9700) = - ((case ((maybe_aq_matches_prefix stringappend_9680)) of - Some (stringappend_9690,stringappend_9700) => - (stringappend_9690, stringappend_9700) + (let stringappend_10040 = (string_drop stringappend_5590 + ((string_length (''sc.'')))) in + (let (aq, stringappend_10060) = + ((case ((maybe_aq_matches_prefix stringappend_10040)) of + Some (stringappend_10050,stringappend_10060) => + (stringappend_10050, stringappend_10060) )) in - (let stringappend_9710 = (string_drop stringappend_9680 - stringappend_9700) in - (let (rl, stringappend_9730) = - ((case ((maybe_rl_matches_prefix stringappend_9710)) of - Some (stringappend_9720,stringappend_9730) => - (stringappend_9720, stringappend_9730) + (let stringappend_10070 = (string_drop stringappend_10040 + stringappend_10060) in + (let (rl, stringappend_10090) = + ((case ((maybe_rl_matches_prefix stringappend_10070)) of + Some (stringappend_10080,stringappend_10090) => + (stringappend_10080, stringappend_10090) )) in - (let stringappend_9740 = (string_drop stringappend_9710 - stringappend_9730) in - (let (size1, stringappend_9760) = + (let stringappend_10100 = (string_drop stringappend_10070 + stringappend_10090) in + (let (size1, stringappend_10120) = ((case ((size_mnemonic_matches_prefix - stringappend_9740)) of - Some (stringappend_9750,stringappend_9760) => - (stringappend_9750, stringappend_9760) + stringappend_10100)) of + Some (stringappend_10110,stringappend_10120) => + (stringappend_10110, stringappend_10120) )) in - (let stringappend_9770 = (string_drop stringappend_9740 - stringappend_9760) in + (let stringappend_10130 = (string_drop stringappend_10100 + stringappend_10120) in (case - (case ((spc_matches_prefix stringappend_9770)) of - Some (stringappend_9780,stringappend_9790) => - (stringappend_9780, stringappend_9790) + (case ((spc_matches_prefix stringappend_10130)) of + Some (stringappend_10140,stringappend_10150) => + (stringappend_10140, stringappend_10150) ) of - (_, stringappend_9790) => - (let stringappend_9800 = (string_drop stringappend_9770 - stringappend_9790) in - (let (rd, stringappend_9820) = - ((case ((reg_name_matches_prefix stringappend_9800 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9810,stringappend_9820) => - (stringappend_9810, stringappend_9820) + (_, stringappend_10150) => + (let stringappend_10160 = (string_drop stringappend_10130 + stringappend_10150) in + (let (rd, stringappend_10180) = + ((case ((reg_name_matches_prefix stringappend_10160 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10170,stringappend_10180) => + (stringappend_10170, stringappend_10180) )) in - (let stringappend_9830 = (string_drop stringappend_9800 - stringappend_9820) in + (let stringappend_10190 = (string_drop stringappend_10160 + stringappend_10180) in (case - (case ((sep_matches_prefix stringappend_9830)) of - Some (stringappend_9840,stringappend_9850) => - (stringappend_9840, stringappend_9850) + (case ((sep_matches_prefix stringappend_10190)) of + Some (stringappend_10200,stringappend_10210) => + (stringappend_10200, stringappend_10210) ) of - (_, stringappend_9850) => - (let stringappend_9860 = (string_drop stringappend_9830 - stringappend_9850) in - (let (rs1, stringappend_9880) = - ((case ((reg_name_matches_prefix stringappend_9860 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9870,stringappend_9880) => - (stringappend_9870, stringappend_9880) + (_, stringappend_10210) => + (let stringappend_10220 = (string_drop stringappend_10190 + stringappend_10210) in + (let (rs1, stringappend_10240) = + ((case ((reg_name_matches_prefix stringappend_10220 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10230,stringappend_10240) => + (stringappend_10230, stringappend_10240) )) in - (let stringappend_9890 = (string_drop stringappend_9860 - stringappend_9880) in + (let stringappend_10250 = (string_drop stringappend_10220 + stringappend_10240) in (case - (case ((sep_matches_prefix stringappend_9890)) of - Some (stringappend_9900,stringappend_9910) => - (stringappend_9900, stringappend_9910) + (case ((sep_matches_prefix stringappend_10250)) of + Some (stringappend_10260,stringappend_10270) => + (stringappend_10260, stringappend_10270) ) of - (_, stringappend_9910) => - (let stringappend_9920 = (string_drop stringappend_9890 - stringappend_9910) in - (let (rs2, stringappend_9940) = - ((case ((reg_name_matches_prefix stringappend_9920 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_9930,stringappend_9940) => - (stringappend_9930, stringappend_9940) + (_, stringappend_10270) => + (let stringappend_10280 = (string_drop stringappend_10250 + stringappend_10270) in + (let (rs2, stringappend_10300) = + ((case ((reg_name_matches_prefix stringappend_10280 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10290,stringappend_10300) => + (stringappend_10290, stringappend_10300) )) in - if(((string_drop stringappend_9920 stringappend_9940)) = + if(((string_drop stringappend_10280 stringappend_10300)) = ('''')) then True else undefined)) )))) )))) )))))))) else - if ((case ((amo_mnemonic_matches_prefix stringappend_5410)) of - Some (stringappend_9960,stringappend_9970) => - (let stringappend_9980 = (string_drop - stringappend_5410 - stringappend_9970) in - if (((((string_startswith stringappend_9980 (''.''))) + if ((case ((amo_mnemonic_matches_prefix stringappend_5590)) of + Some (stringappend_10320,stringappend_10330) => + (let stringappend_10340 = (string_drop + stringappend_5590 + stringappend_10330) in + if (((((string_startswith stringappend_10340 (''.''))) \<and> - ((let stringappend_9990 = (string_drop - stringappend_9980 - ((string_length - (''.'')))) in + ((let stringappend_10350 = (string_drop + stringappend_10340 + ((string_length + (''.'')))) in if ((case ((size_mnemonic_matches_prefix - stringappend_9990)) of - Some (stringappend_10000,stringappend_10010) => - (let stringappend_10020 = - (string_drop stringappend_9990 - stringappend_10010) in + stringappend_10350)) of + Some (stringappend_10360,stringappend_10370) => + (let stringappend_10380 = + (string_drop stringappend_10350 + stringappend_10370) in if ((case ((maybe_aq_matches_prefix - stringappend_10020)) of - Some (stringappend_10030,stringappend_10040) => - (let stringappend_10050 = - (string_drop stringappend_10020 - stringappend_10040) in + stringappend_10380)) of + Some (stringappend_10390,stringappend_10400) => + (let stringappend_10410 = + (string_drop stringappend_10380 + stringappend_10400) in if ((case ((maybe_rl_matches_prefix - stringappend_10050)) of - Some (stringappend_10060,stringappend_10070) => - (let stringappend_10080 = + stringappend_10410)) of + Some (stringappend_10420,stringappend_10430) => + (let stringappend_10440 = (string_drop - stringappend_10050 - stringappend_10070) in + stringappend_10410 + stringappend_10430) in if ((case ((spc_matches_prefix - stringappend_10080)) of - Some (stringappend_10090,stringappend_10100) => - (let stringappend_10110 = + stringappend_10440)) of + Some (stringappend_10450,stringappend_10460) => + (let stringappend_10470 = (string_drop - stringappend_10080 - stringappend_10100) in + stringappend_10440 + stringappend_10460) in if ((case ((reg_name_matches_prefix - stringappend_10110 + stringappend_10470 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10120,stringappend_10130) => - (let stringappend_10140 = + Some (stringappend_10480,stringappend_10490) => + (let stringappend_10500 = (string_drop - stringappend_10110 - stringappend_10130) in + stringappend_10470 + stringappend_10490) in if ((case ((sep_matches_prefix - stringappend_10140)) of - Some (stringappend_10150,stringappend_10160) => - (let stringappend_10170 = + stringappend_10500)) of + Some (stringappend_10510,stringappend_10520) => + (let stringappend_10530 = (string_drop - stringappend_10140 - stringappend_10160) in + stringappend_10500 + stringappend_10520) in if ((case (( reg_name_matches_prefix - stringappend_10170 + stringappend_10530 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10180,stringappend_10190) => - (let stringappend_10200 = + Some (stringappend_10540,stringappend_10550) => + (let stringappend_10560 = (string_drop - stringappend_10170 - stringappend_10190) in + stringappend_10530 + stringappend_10550) in if ((case ((sep_matches_prefix - stringappend_10200)) of + stringappend_10560)) of Some - (stringappend_10210,stringappend_10220) => + (stringappend_10570,stringappend_10580) => (let - stringappend_10230 = + stringappend_10590 = (string_drop - stringappend_10200 - stringappend_10220) in + stringappend_10560 + stringappend_10580) in if ((case ( ( reg_name_matches_prefix - stringappend_10230 + stringappend_10590 :: (( 5 Word.word * ii)) option)) of Some - (stringappend_10240,stringappend_10250) => + (stringappend_10600,stringappend_10610) => if ( ( ( string_drop - stringappend_10230 - stringappend_10250)) + stringappend_10590 + stringappend_10610)) = ('''')) then True else @@ -19352,154 +19602,156 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where True else False) | None => False )) then - (let (op1, stringappend_9970) = + (let (op1, stringappend_10330) = ((case ((amo_mnemonic_matches_prefix - stringappend_5410)) of - Some (stringappend_9960,stringappend_9970) => - (stringappend_9960, stringappend_9970) + stringappend_5590)) of + Some (stringappend_10320,stringappend_10330) => + (stringappend_10320, stringappend_10330) )) in - (let stringappend_9980 = (string_drop stringappend_5410 - stringappend_9970) in - (let stringappend_9990 = (string_drop stringappend_9980 - ((string_length (''.'')))) in - (let (width, stringappend_10010) = + (let stringappend_10340 = (string_drop stringappend_5590 + stringappend_10330) in + (let stringappend_10350 = (string_drop + stringappend_10340 + ((string_length (''.'')))) in + (let (width, stringappend_10370) = ((case ((size_mnemonic_matches_prefix - stringappend_9990)) of - Some (stringappend_10000,stringappend_10010) => - (stringappend_10000, stringappend_10010) + stringappend_10350)) of + Some (stringappend_10360,stringappend_10370) => + (stringappend_10360, stringappend_10370) )) in - (let stringappend_10020 = (string_drop stringappend_9990 - stringappend_10010) in - (let (aq, stringappend_10040) = + (let stringappend_10380 = (string_drop + stringappend_10350 + stringappend_10370) in + (let (aq, stringappend_10400) = ((case ((maybe_aq_matches_prefix - stringappend_10020)) of - Some (stringappend_10030,stringappend_10040) => - (stringappend_10030, stringappend_10040) + stringappend_10380)) of + Some (stringappend_10390,stringappend_10400) => + (stringappend_10390, stringappend_10400) )) in - (let stringappend_10050 = (string_drop - stringappend_10020 - stringappend_10040) in - (let (rl, stringappend_10070) = + (let stringappend_10410 = (string_drop + stringappend_10380 + stringappend_10400) in + (let (rl, stringappend_10430) = ((case ((maybe_rl_matches_prefix - stringappend_10050)) of - Some (stringappend_10060,stringappend_10070) => - (stringappend_10060, stringappend_10070) + stringappend_10410)) of + Some (stringappend_10420,stringappend_10430) => + (stringappend_10420, stringappend_10430) )) in - (let stringappend_10080 = (string_drop - stringappend_10050 - stringappend_10070) in + (let stringappend_10440 = (string_drop + stringappend_10410 + stringappend_10430) in (case - (case ((spc_matches_prefix stringappend_10080)) of - Some (stringappend_10090,stringappend_10100) => - (stringappend_10090, stringappend_10100) + (case ((spc_matches_prefix stringappend_10440)) of + Some (stringappend_10450,stringappend_10460) => + (stringappend_10450, stringappend_10460) ) of - (_, stringappend_10100) => - (let stringappend_10110 = (string_drop - stringappend_10080 - stringappend_10100) in - (let (rd, stringappend_10130) = + (_, stringappend_10460) => + (let stringappend_10470 = (string_drop + stringappend_10440 + stringappend_10460) in + (let (rd, stringappend_10490) = ((case ((reg_name_matches_prefix - stringappend_10110 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10120,stringappend_10130) => - (stringappend_10120, stringappend_10130) + stringappend_10470 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10480,stringappend_10490) => + (stringappend_10480, stringappend_10490) )) in - (let stringappend_10140 = (string_drop - stringappend_10110 - stringappend_10130) in + (let stringappend_10500 = (string_drop + stringappend_10470 + stringappend_10490) in (case - (case ((sep_matches_prefix stringappend_10140)) of - Some (stringappend_10150,stringappend_10160) => - (stringappend_10150, stringappend_10160) + (case ((sep_matches_prefix stringappend_10500)) of + Some (stringappend_10510,stringappend_10520) => + (stringappend_10510, stringappend_10520) ) of - (_, stringappend_10160) => - (let stringappend_10170 = (string_drop - stringappend_10140 - stringappend_10160) in - (let (rs1, stringappend_10190) = + (_, stringappend_10520) => + (let stringappend_10530 = (string_drop + stringappend_10500 + stringappend_10520) in + (let (rs1, stringappend_10550) = ((case ((reg_name_matches_prefix - stringappend_10170 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10180,stringappend_10190) => - (stringappend_10180, stringappend_10190) + stringappend_10530 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10540,stringappend_10550) => + (stringappend_10540, stringappend_10550) )) in - (let stringappend_10200 = (string_drop - stringappend_10170 - stringappend_10190) in + (let stringappend_10560 = (string_drop + stringappend_10530 + stringappend_10550) in (case - (case ((sep_matches_prefix stringappend_10200)) of - Some (stringappend_10210,stringappend_10220) => - (stringappend_10210, stringappend_10220) + (case ((sep_matches_prefix stringappend_10560)) of + Some (stringappend_10570,stringappend_10580) => + (stringappend_10570, stringappend_10580) ) of - (_, stringappend_10220) => - (let stringappend_10230 = (string_drop - stringappend_10200 - stringappend_10220) in - (let (rs2, stringappend_10250) = + (_, stringappend_10580) => + (let stringappend_10590 = (string_drop + stringappend_10560 + stringappend_10580) in + (let (rs2, stringappend_10610) = ((case ((reg_name_matches_prefix - stringappend_10230 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10240,stringappend_10250) => - (stringappend_10240, stringappend_10250) + stringappend_10590 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10600,stringappend_10610) => + (stringappend_10600, stringappend_10610) )) in - if(((string_drop stringappend_10230 stringappend_10250)) + if(((string_drop stringappend_10590 stringappend_10610)) = ('''')) then True else undefined)) )))) )))) )))))))))) else if ((case ((csr_mnemonic_matches_prefix - stringappend_5410)) of - Some (stringappend_10270,stringappend_10280) => - (let stringappend_10290 = (string_drop - stringappend_5410 - stringappend_10280) in - if (((((string_startswith stringappend_10290 (''i''))) + stringappend_5590)) of + Some (stringappend_10630,stringappend_10640) => + (let stringappend_10650 = (string_drop + stringappend_5590 + stringappend_10640) in + if (((((string_startswith stringappend_10650 (''i''))) \<and> - ((let stringappend_10300 = (string_drop - stringappend_10290 + ((let stringappend_10660 = (string_drop + stringappend_10650 ((string_length (''i'')))) in if ((case ((spc_matches_prefix - stringappend_10300)) of - Some (stringappend_10310,stringappend_10320) => - (let stringappend_10330 = - (string_drop stringappend_10300 - stringappend_10320) in + stringappend_10660)) of + Some (stringappend_10670,stringappend_10680) => + (let stringappend_10690 = + (string_drop stringappend_10660 + stringappend_10680) in if ((case ((reg_name_matches_prefix - stringappend_10330 + stringappend_10690 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10340,stringappend_10350) => - (let stringappend_10360 = + Some (stringappend_10700,stringappend_10710) => + (let stringappend_10720 = (string_drop - stringappend_10330 - stringappend_10350) in + stringappend_10690 + stringappend_10710) in if ((case ((sep_matches_prefix - stringappend_10360)) of - Some (stringappend_10370,stringappend_10380) => - (let stringappend_10390 = + stringappend_10720)) of + Some (stringappend_10730,stringappend_10740) => + (let stringappend_10750 = (string_drop - stringappend_10360 - stringappend_10380) in + stringappend_10720 + stringappend_10740) in if ((case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_10390 + stringappend_10750 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10400,stringappend_10410) => - (let stringappend_10420 = + Some (stringappend_10760,stringappend_10770) => + (let stringappend_10780 = (string_drop - stringappend_10390 - stringappend_10410) in + stringappend_10750 + stringappend_10770) in if ((case ((sep_matches_prefix - stringappend_10420)) of - Some (stringappend_10430,stringappend_10440) => - (let stringappend_10450 = + stringappend_10780)) of + Some (stringappend_10790,stringappend_10800) => + (let stringappend_10810 = (string_drop - stringappend_10420 - stringappend_10440) in + stringappend_10780 + stringappend_10800) in if ((case ((csr_name_map_matches_prefix - stringappend_10450 + stringappend_10810 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_10460,stringappend_10470) => + Some (stringappend_10820,stringappend_10830) => if(((string_drop - stringappend_10450 - stringappend_10470)) + stringappend_10810 + stringappend_10830)) = ('''')) then True else False @@ -19521,121 +19773,121 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where True else False) | None => False )) then - (let (op1, stringappend_10280) = + (let (op1, stringappend_10640) = ((case ((csr_mnemonic_matches_prefix - stringappend_5410)) of - Some (stringappend_10270,stringappend_10280) => - (stringappend_10270, stringappend_10280) + stringappend_5590)) of + Some (stringappend_10630,stringappend_10640) => + (stringappend_10630, stringappend_10640) )) in - (let stringappend_10290 = (string_drop - stringappend_5410 - stringappend_10280) in - (let stringappend_10300 = (string_drop - stringappend_10290 + (let stringappend_10650 = (string_drop + stringappend_5590 + stringappend_10640) in + (let stringappend_10660 = (string_drop + stringappend_10650 ((string_length (''i'')))) in (case - (case ((spc_matches_prefix stringappend_10300)) of - Some (stringappend_10310,stringappend_10320) => - (stringappend_10310, stringappend_10320) + (case ((spc_matches_prefix stringappend_10660)) of + Some (stringappend_10670,stringappend_10680) => + (stringappend_10670, stringappend_10680) ) of - (_, stringappend_10320) => - (let stringappend_10330 = (string_drop - stringappend_10300 - stringappend_10320) in - (let (rd, stringappend_10350) = + (_, stringappend_10680) => + (let stringappend_10690 = (string_drop + stringappend_10660 + stringappend_10680) in + (let (rd, stringappend_10710) = ((case ((reg_name_matches_prefix - stringappend_10330 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10340,stringappend_10350) => - (stringappend_10340, stringappend_10350) + stringappend_10690 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10700,stringappend_10710) => + (stringappend_10700, stringappend_10710) )) in - (let stringappend_10360 = (string_drop - stringappend_10330 - stringappend_10350) in + (let stringappend_10720 = (string_drop + stringappend_10690 + stringappend_10710) in (case - (case ((sep_matches_prefix stringappend_10360)) of - Some (stringappend_10370,stringappend_10380) => - (stringappend_10370, stringappend_10380) + (case ((sep_matches_prefix stringappend_10720)) of + Some (stringappend_10730,stringappend_10740) => + (stringappend_10730, stringappend_10740) ) of - (_, stringappend_10380) => - (let stringappend_10390 = (string_drop - stringappend_10360 - stringappend_10380) in - (let (rs1, stringappend_10410) = + (_, stringappend_10740) => + (let stringappend_10750 = (string_drop + stringappend_10720 + stringappend_10740) in + (let (rs1, stringappend_10770) = ((case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_10390 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10400,stringappend_10410) => - (stringappend_10400, stringappend_10410) + stringappend_10750 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10760,stringappend_10770) => + (stringappend_10760, stringappend_10770) )) in - (let stringappend_10420 = (string_drop - stringappend_10390 - stringappend_10410) in + (let stringappend_10780 = (string_drop + stringappend_10750 + stringappend_10770) in (case - (case ((sep_matches_prefix stringappend_10420)) of - Some (stringappend_10430,stringappend_10440) => - (stringappend_10430, stringappend_10440) + (case ((sep_matches_prefix stringappend_10780)) of + Some (stringappend_10790,stringappend_10800) => + (stringappend_10790, stringappend_10800) ) of - (_, stringappend_10440) => - (let stringappend_10450 = (string_drop - stringappend_10420 - stringappend_10440) in - (let (csr, stringappend_10470) = + (_, stringappend_10800) => + (let stringappend_10810 = (string_drop + stringappend_10780 + stringappend_10800) in + (let (csr, stringappend_10830) = ((case ((csr_name_map_matches_prefix - stringappend_10450 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_10460,stringappend_10470) => - (stringappend_10460, stringappend_10470) + stringappend_10810 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_10820,stringappend_10830) => + (stringappend_10820, stringappend_10830) )) in - if(((string_drop stringappend_10450 stringappend_10470)) + if(((string_drop stringappend_10810 stringappend_10830)) = ('''')) then True else undefined)) )))) )))) )))) else if ((case ((csr_mnemonic_matches_prefix - stringappend_5410)) of - Some (stringappend_10490,stringappend_10500) => - (let stringappend_10510 = (string_drop - stringappend_5410 - stringappend_10500) in - if ((case ((spc_matches_prefix stringappend_10510)) of - Some (stringappend_10520,stringappend_10530) => - (let stringappend_10540 = (string_drop - stringappend_10510 - stringappend_10530) in + stringappend_5590)) of + Some (stringappend_10850,stringappend_10860) => + (let stringappend_10870 = (string_drop + stringappend_5590 + stringappend_10860) in + if ((case ((spc_matches_prefix stringappend_10870)) of + Some (stringappend_10880,stringappend_10890) => + (let stringappend_10900 = (string_drop + stringappend_10870 + stringappend_10890) in if ((case ((reg_name_matches_prefix - stringappend_10540 + stringappend_10900 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10550,stringappend_10560) => - (let stringappend_10570 = (string_drop - stringappend_10540 - stringappend_10560) in + Some (stringappend_10910,stringappend_10920) => + (let stringappend_10930 = (string_drop + stringappend_10900 + stringappend_10920) in if ((case ((sep_matches_prefix - stringappend_10570)) of - Some (stringappend_10580,stringappend_10590) => - (let stringappend_10600 = - (string_drop stringappend_10570 - stringappend_10590) in + stringappend_10930)) of + Some (stringappend_10940,stringappend_10950) => + (let stringappend_10960 = + (string_drop stringappend_10930 + stringappend_10950) in if ((case ((reg_name_matches_prefix - stringappend_10600 + stringappend_10960 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10610,stringappend_10620) => - (let stringappend_10630 = + Some (stringappend_10970,stringappend_10980) => + (let stringappend_10990 = (string_drop - stringappend_10600 - stringappend_10620) in + stringappend_10960 + stringappend_10980) in if ((case ((sep_matches_prefix - stringappend_10630)) of - Some (stringappend_10640,stringappend_10650) => - (let stringappend_10660 = + stringappend_10990)) of + Some (stringappend_11000,stringappend_11010) => + (let stringappend_11020 = (string_drop - stringappend_10630 - stringappend_10650) in + stringappend_10990 + stringappend_11010) in if ((case ((csr_name_map_matches_prefix - stringappend_10660 + stringappend_11020 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_10670,stringappend_10680) => + Some (stringappend_11030,stringappend_11040) => if(((string_drop - stringappend_10660 - stringappend_10680)) + stringappend_11020 + stringappend_11040)) = ('''')) then True else False | None => False @@ -19653,118 +19905,118 @@ definition assembly_backwards_matches :: " string \<Rightarrow> bool " where )) then True else False) | None => False )) then - (let (op1, stringappend_10500) = + (let (op1, stringappend_10860) = ((case ((csr_mnemonic_matches_prefix - stringappend_5410)) of - Some (stringappend_10490,stringappend_10500) => - (stringappend_10490, stringappend_10500) + stringappend_5590)) of + Some (stringappend_10850,stringappend_10860) => + (stringappend_10850, stringappend_10860) )) in - (let stringappend_10510 = (string_drop - stringappend_5410 - stringappend_10500) in + (let stringappend_10870 = (string_drop + stringappend_5590 + stringappend_10860) in (case - (case ((spc_matches_prefix stringappend_10510)) of - Some (stringappend_10520,stringappend_10530) => - (stringappend_10520, stringappend_10530) + (case ((spc_matches_prefix stringappend_10870)) of + Some (stringappend_10880,stringappend_10890) => + (stringappend_10880, stringappend_10890) ) of - (_, stringappend_10530) => - (let stringappend_10540 = (string_drop - stringappend_10510 - stringappend_10530) in - (let (rd, stringappend_10560) = + (_, stringappend_10890) => + (let stringappend_10900 = (string_drop + stringappend_10870 + stringappend_10890) in + (let (rd, stringappend_10920) = ((case ((reg_name_matches_prefix - stringappend_10540 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10550,stringappend_10560) => - (stringappend_10550, stringappend_10560) + stringappend_10900 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10910,stringappend_10920) => + (stringappend_10910, stringappend_10920) )) in - (let stringappend_10570 = (string_drop - stringappend_10540 - stringappend_10560) in + (let stringappend_10930 = (string_drop + stringappend_10900 + stringappend_10920) in (case - (case ((sep_matches_prefix stringappend_10570)) of - Some (stringappend_10580,stringappend_10590) => - (stringappend_10580, stringappend_10590) + (case ((sep_matches_prefix stringappend_10930)) of + Some (stringappend_10940,stringappend_10950) => + (stringappend_10940, stringappend_10950) ) of - (_, stringappend_10590) => - (let stringappend_10600 = (string_drop - stringappend_10570 - stringappend_10590) in - (let (rs1, stringappend_10620) = + (_, stringappend_10950) => + (let stringappend_10960 = (string_drop + stringappend_10930 + stringappend_10950) in + (let (rs1, stringappend_10980) = ((case ((reg_name_matches_prefix - stringappend_10600 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_10610,stringappend_10620) => - (stringappend_10610, stringappend_10620) + stringappend_10960 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_10970,stringappend_10980) => + (stringappend_10970, stringappend_10980) )) in - (let stringappend_10630 = (string_drop - stringappend_10600 - stringappend_10620) in + (let stringappend_10990 = (string_drop + stringappend_10960 + stringappend_10980) in (case - (case ((sep_matches_prefix stringappend_10630)) of - Some (stringappend_10640,stringappend_10650) => - (stringappend_10640, stringappend_10650) + (case ((sep_matches_prefix stringappend_10990)) of + Some (stringappend_11000,stringappend_11010) => + (stringappend_11000, stringappend_11010) ) of - (_, stringappend_10650) => - (let stringappend_10660 = (string_drop - stringappend_10630 - stringappend_10650) in - (let (csr, stringappend_10680) = + (_, stringappend_11010) => + (let stringappend_11020 = (string_drop + stringappend_10990 + stringappend_11010) in + (let (csr, stringappend_11040) = ((case ((csr_name_map_matches_prefix - stringappend_10660 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_10670,stringappend_10680) => - (stringappend_10670, stringappend_10680) + stringappend_11020 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_11030,stringappend_11040) => + (stringappend_11030, stringappend_11040) )) in - if(((string_drop stringappend_10660 - stringappend_10680)) = ('''')) then True else + if(((string_drop stringappend_11020 + stringappend_11040)) = ('''')) then True else undefined)) )))) )))) ))) else - if (((((string_startswith stringappend_5410 + if (((((string_startswith stringappend_5590 (''illegal''))) \<and> - ((let stringappend_10700 = - (string_drop stringappend_5410 + ((let stringappend_11060 = + (string_drop stringappend_5590 ((string_length (''illegal'')))) in if ((case ((spc_matches_prefix - stringappend_10700)) of - Some (stringappend_10710,stringappend_10720) => - (let stringappend_10730 = - (string_drop stringappend_10700 - stringappend_10720) in + stringappend_11060)) of + Some (stringappend_11070,stringappend_11080) => + (let stringappend_11090 = + (string_drop stringappend_11060 + stringappend_11080) in if ((case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_10730 + stringappend_11090 :: (( 32 Word.word * ii)) option)) of - Some (stringappend_10740,stringappend_10750) => - if(((string_drop stringappend_10730 - stringappend_10750)) = + Some (stringappend_11100,stringappend_11110) => + if(((string_drop stringappend_11090 + stringappend_11110)) = ('''')) then True else False | None => False )) then True else False) | None => False )) then True else False))))) then - (let stringappend_10700 = (string_drop - stringappend_5410 + (let stringappend_11060 = (string_drop + stringappend_5590 ((string_length (''illegal'')))) in (case - (case ((spc_matches_prefix stringappend_10700)) of - Some (stringappend_10710,stringappend_10720) => - (stringappend_10710, stringappend_10720) + (case ((spc_matches_prefix stringappend_11060)) of + Some (stringappend_11070,stringappend_11080) => + (stringappend_11070, stringappend_11080) ) of - (_, stringappend_10720) => - (let stringappend_10730 = (string_drop - stringappend_10700 - stringappend_10720) in - (let (s, stringappend_10750) = + (_, stringappend_11080) => + (let stringappend_11090 = (string_drop + stringappend_11060 + stringappend_11080) in + (let (s, stringappend_11110) = ((case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_10730 :: (( 32 Word.word * ii)) option)) of - Some (stringappend_10740,stringappend_10750) => - (stringappend_10740, stringappend_10750) + stringappend_11090 :: (( 32 Word.word * ii)) option)) of + Some (stringappend_11100,stringappend_11110) => + (stringappend_11100, stringappend_11110) )) in - if(((string_drop stringappend_10730 - stringappend_10750)) = ('''')) then + if(((string_drop stringappend_11090 + stringappend_11110)) = ('''')) then True else undefined)) )) else False))))))))" @@ -20957,7 +21209,7 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) ))) - else if ((case ((mul_mnemonic_matches_prefix stringappend_00)) of + else if ((case ((shiftiwop_mnemonic_matches_prefix stringappend_00)) of Some (stringappend_2490,stringappend_2500) => (let stringappend_2510 = (string_drop stringappend_00 stringappend_2500) in if ((case ((spc_matches_prefix stringappend_2510)) of @@ -20974,18 +21226,11 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " Some (stringappend_2610,stringappend_2620) => (let stringappend_2630 = (string_drop stringappend_2600 stringappend_2620) in - if ((case ((sep_matches_prefix stringappend_2630)) of - Some (stringappend_2640,stringappend_2650) => - (let stringappend_2660 = - (string_drop stringappend_2630 stringappend_2650) in - if ((case ((reg_name_matches_prefix stringappend_2660 - :: (( 5 Word.word * ii))option)) of - Some (stringappend_2670,stringappend_2680) => - (case ((string_drop stringappend_2660 stringappend_2680)) of s0 => True ) - | None => False - )) then - True - else False) + if ((case ((hex_bits_5_matches_prefix + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_2630 + :: (( 5 Word.word * ii))option)) of + Some (stringappend_2640,stringappend_2650) => + (case ((string_drop stringappend_2630 stringappend_2650)) of s0 => True ) | None => False )) then True @@ -21008,8 +21253,8 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " else False) | None => False )) then - (let ((high, signed1, signed2), stringappend_2500) = - ((case ((mul_mnemonic_matches_prefix stringappend_00)) of + (let (op1, stringappend_2500) = + ((case ((shiftiwop_mnemonic_matches_prefix stringappend_00)) of Some (stringappend_2490,stringappend_2500) => (stringappend_2490, stringappend_2500) )) in (let stringappend_2510 = (string_drop stringappend_00 stringappend_2500) in @@ -21035,17 +21280,109 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " Some (stringappend_2610,stringappend_2620) => (stringappend_2610, stringappend_2620) )) in (let stringappend_2630 = (string_drop stringappend_2600 stringappend_2620) in + (let (shamt, stringappend_2650) = + ((case ((hex_bits_5_matches_prefix + instance_Sail2_values_Bitvector_Machine_word_mword_dict + stringappend_2630 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_2640,stringappend_2650) => (stringappend_2640, stringappend_2650) + )) in + (case ((string_drop stringappend_2630 stringappend_2650)) of + s0 => + Some + (SHIFTIWOP (shamt,rs1,rd,op1), ((string_length arg0)) - + ((string_length s0))) + ))))) + )))) + ))) + else if ((case ((mul_mnemonic_matches_prefix stringappend_00)) of + Some (stringappend_2670,stringappend_2680) => + (let stringappend_2690 = (string_drop stringappend_00 stringappend_2680) in + if ((case ((spc_matches_prefix stringappend_2690)) of + Some (stringappend_2700,stringappend_2710) => + (let stringappend_2720 = (string_drop stringappend_2690 stringappend_2710) in + if ((case ((reg_name_matches_prefix stringappend_2720 :: (( 5 Word.word * ii))option)) of + Some (stringappend_2730,stringappend_2740) => + (let stringappend_2750 = (string_drop stringappend_2720 stringappend_2740) in + if ((case ((sep_matches_prefix stringappend_2750)) of + Some (stringappend_2760,stringappend_2770) => + (let stringappend_2780 = (string_drop stringappend_2750 stringappend_2770) in + if ((case ((reg_name_matches_prefix stringappend_2780 + :: (( 5 Word.word * ii))option)) of + Some (stringappend_2790,stringappend_2800) => + (let stringappend_2810 = + (string_drop stringappend_2780 stringappend_2800) in + if ((case ((sep_matches_prefix stringappend_2810)) of + Some (stringappend_2820,stringappend_2830) => + (let stringappend_2840 = + (string_drop stringappend_2810 stringappend_2830) in + if ((case ((reg_name_matches_prefix stringappend_2840 + :: (( 5 Word.word * ii))option)) of + Some (stringappend_2850,stringappend_2860) => + (case ((string_drop stringappend_2840 stringappend_2860)) of s0 => True ) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + True + else False) + | None => False + )) then + (let ((high, signed1, signed2), stringappend_2680) = + ((case ((mul_mnemonic_matches_prefix stringappend_00)) of + Some (stringappend_2670,stringappend_2680) => (stringappend_2670, stringappend_2680) + )) in + (let stringappend_2690 = (string_drop stringappend_00 stringappend_2680) in + (case + (case ((spc_matches_prefix stringappend_2690)) of + Some (stringappend_2700,stringappend_2710) => (stringappend_2700, stringappend_2710) + ) of + (_, stringappend_2710) => + (let stringappend_2720 = (string_drop stringappend_2690 stringappend_2710) in + (let (rd, stringappend_2740) = + ((case ((reg_name_matches_prefix stringappend_2720 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_2730,stringappend_2740) => (stringappend_2730, stringappend_2740) + )) in + (let stringappend_2750 = (string_drop stringappend_2720 stringappend_2740) in (case - (case ((sep_matches_prefix stringappend_2630)) of - Some (stringappend_2640,stringappend_2650) => (stringappend_2640, stringappend_2650) + (case ((sep_matches_prefix stringappend_2750)) of + Some (stringappend_2760,stringappend_2770) => (stringappend_2760, stringappend_2770) ) of - (_, stringappend_2650) => - (let stringappend_2660 = (string_drop stringappend_2630 stringappend_2650) in - (let (rs2, stringappend_2680) = - ((case ((reg_name_matches_prefix stringappend_2660 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_2670,stringappend_2680) => (stringappend_2670, stringappend_2680) + (_, stringappend_2770) => + (let stringappend_2780 = (string_drop stringappend_2750 stringappend_2770) in + (let (rs1, stringappend_2800) = + ((case ((reg_name_matches_prefix stringappend_2780 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_2790,stringappend_2800) => (stringappend_2790, stringappend_2800) )) in - (case ((string_drop stringappend_2660 stringappend_2680)) of + (let stringappend_2810 = (string_drop stringappend_2780 stringappend_2800) in + (case + (case ((sep_matches_prefix stringappend_2810)) of + Some (stringappend_2820,stringappend_2830) => (stringappend_2820, stringappend_2830) + ) of + (_, stringappend_2830) => + (let stringappend_2840 = (string_drop stringappend_2810 stringappend_2830) in + (let (rs2, stringappend_2860) = + ((case ((reg_name_matches_prefix stringappend_2840 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_2850,stringappend_2860) => (stringappend_2850, stringappend_2860) + )) in + (case ((string_drop stringappend_2840 stringappend_2860)) of s0 => Some (MUL (rs2,rs1,rd,high,signed1,signed2), @@ -21054,35 +21391,35 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) ))) - else if (((((string_startswith stringappend_00 (''div''))) \<and> ((let stringappend_2700 = (string_drop stringappend_00 ((string_length (''div'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_2700)) of - Some (stringappend_2710,stringappend_2720) => - (let stringappend_2730 = (string_drop stringappend_2700 stringappend_2720) in - if ((case ((spc_matches_prefix stringappend_2730)) of - Some (stringappend_2740,stringappend_2750) => - (let stringappend_2760 = (string_drop stringappend_2730 stringappend_2750) in - if ((case ((reg_name_matches_prefix stringappend_2760 + else if (((((string_startswith stringappend_00 (''div''))) \<and> ((let stringappend_2880 = (string_drop stringappend_00 ((string_length (''div'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_2880)) of + Some (stringappend_2890,stringappend_2900) => + (let stringappend_2910 = (string_drop stringappend_2880 stringappend_2900) in + if ((case ((spc_matches_prefix stringappend_2910)) of + Some (stringappend_2920,stringappend_2930) => + (let stringappend_2940 = (string_drop stringappend_2910 stringappend_2930) in + if ((case ((reg_name_matches_prefix stringappend_2940 :: (( 5 Word.word * ii))option)) of - Some (stringappend_2770,stringappend_2780) => - (let stringappend_2790 = - (string_drop stringappend_2760 stringappend_2780) in - if ((case ((sep_matches_prefix stringappend_2790)) of - Some (stringappend_2800,stringappend_2810) => - (let stringappend_2820 = - (string_drop stringappend_2790 stringappend_2810) in - if ((case ((reg_name_matches_prefix stringappend_2820 + Some (stringappend_2950,stringappend_2960) => + (let stringappend_2970 = + (string_drop stringappend_2940 stringappend_2960) in + if ((case ((sep_matches_prefix stringappend_2970)) of + Some (stringappend_2980,stringappend_2990) => + (let stringappend_3000 = + (string_drop stringappend_2970 stringappend_2990) in + if ((case ((reg_name_matches_prefix stringappend_3000 :: (( 5 Word.word * ii))option)) of - Some (stringappend_2830,stringappend_2840) => - (let stringappend_2850 = - (string_drop stringappend_2820 stringappend_2840) in - if ((case ((sep_matches_prefix stringappend_2850)) of - Some (stringappend_2860,stringappend_2870) => - (let stringappend_2880 = - (string_drop stringappend_2850 stringappend_2870) in - if ((case ((reg_name_matches_prefix stringappend_2880 + Some (stringappend_3010,stringappend_3020) => + (let stringappend_3030 = + (string_drop stringappend_3000 stringappend_3020) in + if ((case ((sep_matches_prefix stringappend_3030)) of + Some (stringappend_3040,stringappend_3050) => + (let stringappend_3060 = + (string_drop stringappend_3030 stringappend_3050) in + if ((case ((reg_name_matches_prefix stringappend_3060 :: (( 5 Word.word * ii))option)) of - Some (stringappend_2890,stringappend_2900) => - (case ((string_drop stringappend_2880 stringappend_2900)) of s0 => True ) + Some (stringappend_3070,stringappend_3080) => + (case ((string_drop stringappend_3060 stringappend_3080)) of s0 => True ) | None => False )) then True @@ -21111,45 +21448,45 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_2700 = (string_drop stringappend_00 ((string_length (''div'')))) in - (let (s, stringappend_2720) = - ((case ((maybe_not_u_matches_prefix stringappend_2700)) of - Some (stringappend_2710,stringappend_2720) => (stringappend_2710, stringappend_2720) + (let stringappend_2880 = (string_drop stringappend_00 ((string_length (''div'')))) in + (let (s, stringappend_2900) = + ((case ((maybe_not_u_matches_prefix stringappend_2880)) of + Some (stringappend_2890,stringappend_2900) => (stringappend_2890, stringappend_2900) )) in - (let stringappend_2730 = (string_drop stringappend_2700 stringappend_2720) in + (let stringappend_2910 = (string_drop stringappend_2880 stringappend_2900) in (case - (case ((spc_matches_prefix stringappend_2730)) of - Some (stringappend_2740,stringappend_2750) => (stringappend_2740, stringappend_2750) + (case ((spc_matches_prefix stringappend_2910)) of + Some (stringappend_2920,stringappend_2930) => (stringappend_2920, stringappend_2930) ) of - (_, stringappend_2750) => - (let stringappend_2760 = (string_drop stringappend_2730 stringappend_2750) in - (let (rd, stringappend_2780) = - ((case ((reg_name_matches_prefix stringappend_2760 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_2770,stringappend_2780) => (stringappend_2770, stringappend_2780) + (_, stringappend_2930) => + (let stringappend_2940 = (string_drop stringappend_2910 stringappend_2930) in + (let (rd, stringappend_2960) = + ((case ((reg_name_matches_prefix stringappend_2940 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_2950,stringappend_2960) => (stringappend_2950, stringappend_2960) )) in - (let stringappend_2790 = (string_drop stringappend_2760 stringappend_2780) in + (let stringappend_2970 = (string_drop stringappend_2940 stringappend_2960) in (case - (case ((sep_matches_prefix stringappend_2790)) of - Some (stringappend_2800,stringappend_2810) => (stringappend_2800, stringappend_2810) + (case ((sep_matches_prefix stringappend_2970)) of + Some (stringappend_2980,stringappend_2990) => (stringappend_2980, stringappend_2990) ) of - (_, stringappend_2810) => - (let stringappend_2820 = (string_drop stringappend_2790 stringappend_2810) in - (let (rs1, stringappend_2840) = - ((case ((reg_name_matches_prefix stringappend_2820 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_2830,stringappend_2840) => (stringappend_2830, stringappend_2840) + (_, stringappend_2990) => + (let stringappend_3000 = (string_drop stringappend_2970 stringappend_2990) in + (let (rs1, stringappend_3020) = + ((case ((reg_name_matches_prefix stringappend_3000 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3010,stringappend_3020) => (stringappend_3010, stringappend_3020) )) in - (let stringappend_2850 = (string_drop stringappend_2820 stringappend_2840) in + (let stringappend_3030 = (string_drop stringappend_3000 stringappend_3020) in (case - (case ((sep_matches_prefix stringappend_2850)) of - Some (stringappend_2860,stringappend_2870) => (stringappend_2860, stringappend_2870) + (case ((sep_matches_prefix stringappend_3030)) of + Some (stringappend_3040,stringappend_3050) => (stringappend_3040, stringappend_3050) ) of - (_, stringappend_2870) => - (let stringappend_2880 = (string_drop stringappend_2850 stringappend_2870) in - (let (rs2, stringappend_2900) = - ((case ((reg_name_matches_prefix stringappend_2880 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_2890,stringappend_2900) => (stringappend_2890, stringappend_2900) + (_, stringappend_3050) => + (let stringappend_3060 = (string_drop stringappend_3030 stringappend_3050) in + (let (rs2, stringappend_3080) = + ((case ((reg_name_matches_prefix stringappend_3060 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3070,stringappend_3080) => (stringappend_3070, stringappend_3080) )) in - (case ((string_drop stringappend_2880 stringappend_2900)) of + (case ((string_drop stringappend_3060 stringappend_3080)) of s1 => Some (DIV (rs2,rs1,rd,s), ((string_length arg0)) - ((string_length s1))) @@ -21157,35 +21494,35 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) )))) - else if (((((string_startswith stringappend_00 (''rem''))) \<and> ((let stringappend_2920 = (string_drop stringappend_00 ((string_length (''rem'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_2920)) of - Some (stringappend_2930,stringappend_2940) => - (let stringappend_2950 = (string_drop stringappend_2920 stringappend_2940) in - if ((case ((spc_matches_prefix stringappend_2950)) of - Some (stringappend_2960,stringappend_2970) => - (let stringappend_2980 = (string_drop stringappend_2950 stringappend_2970) in - if ((case ((reg_name_matches_prefix stringappend_2980 + else if (((((string_startswith stringappend_00 (''rem''))) \<and> ((let stringappend_3100 = (string_drop stringappend_00 ((string_length (''rem'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_3100)) of + Some (stringappend_3110,stringappend_3120) => + (let stringappend_3130 = (string_drop stringappend_3100 stringappend_3120) in + if ((case ((spc_matches_prefix stringappend_3130)) of + Some (stringappend_3140,stringappend_3150) => + (let stringappend_3160 = (string_drop stringappend_3130 stringappend_3150) in + if ((case ((reg_name_matches_prefix stringappend_3160 :: (( 5 Word.word * ii))option)) of - Some (stringappend_2990,stringappend_3000) => - (let stringappend_3010 = - (string_drop stringappend_2980 stringappend_3000) in - if ((case ((sep_matches_prefix stringappend_3010)) of - Some (stringappend_3020,stringappend_3030) => - (let stringappend_3040 = - (string_drop stringappend_3010 stringappend_3030) in - if ((case ((reg_name_matches_prefix stringappend_3040 + Some (stringappend_3170,stringappend_3180) => + (let stringappend_3190 = + (string_drop stringappend_3160 stringappend_3180) in + if ((case ((sep_matches_prefix stringappend_3190)) of + Some (stringappend_3200,stringappend_3210) => + (let stringappend_3220 = + (string_drop stringappend_3190 stringappend_3210) in + if ((case ((reg_name_matches_prefix stringappend_3220 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3050,stringappend_3060) => - (let stringappend_3070 = - (string_drop stringappend_3040 stringappend_3060) in - if ((case ((sep_matches_prefix stringappend_3070)) of - Some (stringappend_3080,stringappend_3090) => - (let stringappend_3100 = - (string_drop stringappend_3070 stringappend_3090) in - if ((case ((reg_name_matches_prefix stringappend_3100 + Some (stringappend_3230,stringappend_3240) => + (let stringappend_3250 = + (string_drop stringappend_3220 stringappend_3240) in + if ((case ((sep_matches_prefix stringappend_3250)) of + Some (stringappend_3260,stringappend_3270) => + (let stringappend_3280 = + (string_drop stringappend_3250 stringappend_3270) in + if ((case ((reg_name_matches_prefix stringappend_3280 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3110,stringappend_3120) => - (case ((string_drop stringappend_3100 stringappend_3120)) of s0 => True ) + Some (stringappend_3290,stringappend_3300) => + (case ((string_drop stringappend_3280 stringappend_3300)) of s0 => True ) | None => False )) then True @@ -21214,45 +21551,45 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_2920 = (string_drop stringappend_00 ((string_length (''rem'')))) in - (let (s, stringappend_2940) = - ((case ((maybe_not_u_matches_prefix stringappend_2920)) of - Some (stringappend_2930,stringappend_2940) => (stringappend_2930, stringappend_2940) + (let stringappend_3100 = (string_drop stringappend_00 ((string_length (''rem'')))) in + (let (s, stringappend_3120) = + ((case ((maybe_not_u_matches_prefix stringappend_3100)) of + Some (stringappend_3110,stringappend_3120) => (stringappend_3110, stringappend_3120) )) in - (let stringappend_2950 = (string_drop stringappend_2920 stringappend_2940) in + (let stringappend_3130 = (string_drop stringappend_3100 stringappend_3120) in (case - (case ((spc_matches_prefix stringappend_2950)) of - Some (stringappend_2960,stringappend_2970) => (stringappend_2960, stringappend_2970) + (case ((spc_matches_prefix stringappend_3130)) of + Some (stringappend_3140,stringappend_3150) => (stringappend_3140, stringappend_3150) ) of - (_, stringappend_2970) => - (let stringappend_2980 = (string_drop stringappend_2950 stringappend_2970) in - (let (rd, stringappend_3000) = - ((case ((reg_name_matches_prefix stringappend_2980 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_2990,stringappend_3000) => (stringappend_2990, stringappend_3000) + (_, stringappend_3150) => + (let stringappend_3160 = (string_drop stringappend_3130 stringappend_3150) in + (let (rd, stringappend_3180) = + ((case ((reg_name_matches_prefix stringappend_3160 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3170,stringappend_3180) => (stringappend_3170, stringappend_3180) )) in - (let stringappend_3010 = (string_drop stringappend_2980 stringappend_3000) in + (let stringappend_3190 = (string_drop stringappend_3160 stringappend_3180) in (case - (case ((sep_matches_prefix stringappend_3010)) of - Some (stringappend_3020,stringappend_3030) => (stringappend_3020, stringappend_3030) + (case ((sep_matches_prefix stringappend_3190)) of + Some (stringappend_3200,stringappend_3210) => (stringappend_3200, stringappend_3210) ) of - (_, stringappend_3030) => - (let stringappend_3040 = (string_drop stringappend_3010 stringappend_3030) in - (let (rs1, stringappend_3060) = - ((case ((reg_name_matches_prefix stringappend_3040 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3050,stringappend_3060) => (stringappend_3050, stringappend_3060) + (_, stringappend_3210) => + (let stringappend_3220 = (string_drop stringappend_3190 stringappend_3210) in + (let (rs1, stringappend_3240) = + ((case ((reg_name_matches_prefix stringappend_3220 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3230,stringappend_3240) => (stringappend_3230, stringappend_3240) )) in - (let stringappend_3070 = (string_drop stringappend_3040 stringappend_3060) in + (let stringappend_3250 = (string_drop stringappend_3220 stringappend_3240) in (case - (case ((sep_matches_prefix stringappend_3070)) of - Some (stringappend_3080,stringappend_3090) => (stringappend_3080, stringappend_3090) + (case ((sep_matches_prefix stringappend_3250)) of + Some (stringappend_3260,stringappend_3270) => (stringappend_3260, stringappend_3270) ) of - (_, stringappend_3090) => - (let stringappend_3100 = (string_drop stringappend_3070 stringappend_3090) in - (let (rs2, stringappend_3120) = - ((case ((reg_name_matches_prefix stringappend_3100 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3110,stringappend_3120) => (stringappend_3110, stringappend_3120) + (_, stringappend_3270) => + (let stringappend_3280 = (string_drop stringappend_3250 stringappend_3270) in + (let (rs2, stringappend_3300) = + ((case ((reg_name_matches_prefix stringappend_3280 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3290,stringappend_3300) => (stringappend_3290, stringappend_3300) )) in - (case ((string_drop stringappend_3100 stringappend_3120)) of + (case ((string_drop stringappend_3280 stringappend_3300)) of s1 => Some (REM (rs2,rs1,rd,s), ((string_length arg0)) - ((string_length s1))) @@ -21260,31 +21597,31 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) )))) - else if (((((string_startswith stringappend_00 (''mulw''))) \<and> ((let stringappend_3140 = (string_drop stringappend_00 ((string_length (''mulw'')))) in - if ((case ((spc_matches_prefix stringappend_3140)) of - Some (stringappend_3150,stringappend_3160) => - (let stringappend_3170 = (string_drop stringappend_3140 stringappend_3160) in - if ((case ((reg_name_matches_prefix stringappend_3170 + else if (((((string_startswith stringappend_00 (''mulw''))) \<and> ((let stringappend_3320 = (string_drop stringappend_00 ((string_length (''mulw'')))) in + if ((case ((spc_matches_prefix stringappend_3320)) of + Some (stringappend_3330,stringappend_3340) => + (let stringappend_3350 = (string_drop stringappend_3320 stringappend_3340) in + if ((case ((reg_name_matches_prefix stringappend_3350 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3180,stringappend_3190) => - (let stringappend_3200 = (string_drop stringappend_3170 stringappend_3190) in - if ((case ((sep_matches_prefix stringappend_3200)) of - Some (stringappend_3210,stringappend_3220) => - (let stringappend_3230 = - (string_drop stringappend_3200 stringappend_3220) in - if ((case ((reg_name_matches_prefix stringappend_3230 + Some (stringappend_3360,stringappend_3370) => + (let stringappend_3380 = (string_drop stringappend_3350 stringappend_3370) in + if ((case ((sep_matches_prefix stringappend_3380)) of + Some (stringappend_3390,stringappend_3400) => + (let stringappend_3410 = + (string_drop stringappend_3380 stringappend_3400) in + if ((case ((reg_name_matches_prefix stringappend_3410 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3240,stringappend_3250) => - (let stringappend_3260 = - (string_drop stringappend_3230 stringappend_3250) in - if ((case ((sep_matches_prefix stringappend_3260)) of - Some (stringappend_3270,stringappend_3280) => - (let stringappend_3290 = - (string_drop stringappend_3260 stringappend_3280) in - if ((case ((reg_name_matches_prefix stringappend_3290 + Some (stringappend_3420,stringappend_3430) => + (let stringappend_3440 = + (string_drop stringappend_3410 stringappend_3430) in + if ((case ((sep_matches_prefix stringappend_3440)) of + Some (stringappend_3450,stringappend_3460) => + (let stringappend_3470 = + (string_drop stringappend_3440 stringappend_3460) in + if ((case ((reg_name_matches_prefix stringappend_3470 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3300,stringappend_3310) => - (case ((string_drop stringappend_3290 stringappend_3310)) of s0 => True ) + Some (stringappend_3480,stringappend_3490) => + (case ((string_drop stringappend_3470 stringappend_3490)) of s0 => True ) | None => False )) then True @@ -21309,40 +21646,40 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_3140 = (string_drop stringappend_00 ((string_length (''mulw'')))) in + (let stringappend_3320 = (string_drop stringappend_00 ((string_length (''mulw'')))) in (case - (case ((spc_matches_prefix stringappend_3140)) of - Some (stringappend_3150,stringappend_3160) => (stringappend_3150, stringappend_3160) + (case ((spc_matches_prefix stringappend_3320)) of + Some (stringappend_3330,stringappend_3340) => (stringappend_3330, stringappend_3340) ) of - (_, stringappend_3160) => - (let stringappend_3170 = (string_drop stringappend_3140 stringappend_3160) in - (let (rd, stringappend_3190) = - ((case ((reg_name_matches_prefix stringappend_3170 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3180,stringappend_3190) => (stringappend_3180, stringappend_3190) + (_, stringappend_3340) => + (let stringappend_3350 = (string_drop stringappend_3320 stringappend_3340) in + (let (rd, stringappend_3370) = + ((case ((reg_name_matches_prefix stringappend_3350 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3360,stringappend_3370) => (stringappend_3360, stringappend_3370) )) in - (let stringappend_3200 = (string_drop stringappend_3170 stringappend_3190) in + (let stringappend_3380 = (string_drop stringappend_3350 stringappend_3370) in (case - (case ((sep_matches_prefix stringappend_3200)) of - Some (stringappend_3210,stringappend_3220) => (stringappend_3210, stringappend_3220) + (case ((sep_matches_prefix stringappend_3380)) of + Some (stringappend_3390,stringappend_3400) => (stringappend_3390, stringappend_3400) ) of - (_, stringappend_3220) => - (let stringappend_3230 = (string_drop stringappend_3200 stringappend_3220) in - (let (rs1, stringappend_3250) = - ((case ((reg_name_matches_prefix stringappend_3230 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3240,stringappend_3250) => (stringappend_3240, stringappend_3250) + (_, stringappend_3400) => + (let stringappend_3410 = (string_drop stringappend_3380 stringappend_3400) in + (let (rs1, stringappend_3430) = + ((case ((reg_name_matches_prefix stringappend_3410 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3420,stringappend_3430) => (stringappend_3420, stringappend_3430) )) in - (let stringappend_3260 = (string_drop stringappend_3230 stringappend_3250) in + (let stringappend_3440 = (string_drop stringappend_3410 stringappend_3430) in (case - (case ((sep_matches_prefix stringappend_3260)) of - Some (stringappend_3270,stringappend_3280) => (stringappend_3270, stringappend_3280) + (case ((sep_matches_prefix stringappend_3440)) of + Some (stringappend_3450,stringappend_3460) => (stringappend_3450, stringappend_3460) ) of - (_, stringappend_3280) => - (let stringappend_3290 = (string_drop stringappend_3260 stringappend_3280) in - (let (rs2, stringappend_3310) = - ((case ((reg_name_matches_prefix stringappend_3290 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3300,stringappend_3310) => (stringappend_3300, stringappend_3310) + (_, stringappend_3460) => + (let stringappend_3470 = (string_drop stringappend_3440 stringappend_3460) in + (let (rs2, stringappend_3490) = + ((case ((reg_name_matches_prefix stringappend_3470 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3480,stringappend_3490) => (stringappend_3480, stringappend_3490) )) in - (case ((string_drop stringappend_3290 stringappend_3310)) of + (case ((string_drop stringappend_3470 stringappend_3490)) of s0 => Some (MULW (rs2,rs1,rd), ((string_length arg0)) - ((string_length s0))) @@ -21350,40 +21687,40 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) )) - else if (((((string_startswith stringappend_00 (''div''))) \<and> ((let stringappend_3330 = (string_drop stringappend_00 ((string_length (''div'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_3330)) of - Some (stringappend_3340,stringappend_3350) => - (let stringappend_3360 = (string_drop stringappend_3330 stringappend_3350) in - if (((((string_startswith stringappend_3360 (''w''))) \<and> ((let stringappend_3370 = - (string_drop stringappend_3360 ((string_length (''w'')))) in - if ((case ((spc_matches_prefix stringappend_3370)) of - Some (stringappend_3380,stringappend_3390) => - (let stringappend_3400 = - (string_drop stringappend_3370 stringappend_3390) in - if ((case ((reg_name_matches_prefix stringappend_3400 + else if (((((string_startswith stringappend_00 (''div''))) \<and> ((let stringappend_3510 = (string_drop stringappend_00 ((string_length (''div'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_3510)) of + Some (stringappend_3520,stringappend_3530) => + (let stringappend_3540 = (string_drop stringappend_3510 stringappend_3530) in + if (((((string_startswith stringappend_3540 (''w''))) \<and> ((let stringappend_3550 = + (string_drop stringappend_3540 ((string_length (''w'')))) in + if ((case ((spc_matches_prefix stringappend_3550)) of + Some (stringappend_3560,stringappend_3570) => + (let stringappend_3580 = + (string_drop stringappend_3550 stringappend_3570) in + if ((case ((reg_name_matches_prefix stringappend_3580 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3410,stringappend_3420) => - (let stringappend_3430 = - (string_drop stringappend_3400 stringappend_3420) in - if ((case ((sep_matches_prefix stringappend_3430)) of - Some (stringappend_3440,stringappend_3450) => - (let stringappend_3460 = - (string_drop stringappend_3430 stringappend_3450) in - if ((case ((reg_name_matches_prefix stringappend_3460 + Some (stringappend_3590,stringappend_3600) => + (let stringappend_3610 = + (string_drop stringappend_3580 stringappend_3600) in + if ((case ((sep_matches_prefix stringappend_3610)) of + Some (stringappend_3620,stringappend_3630) => + (let stringappend_3640 = + (string_drop stringappend_3610 stringappend_3630) in + if ((case ((reg_name_matches_prefix stringappend_3640 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3470,stringappend_3480) => - (let stringappend_3490 = - (string_drop stringappend_3460 stringappend_3480) in - if ((case ((sep_matches_prefix stringappend_3490)) of - Some (stringappend_3500,stringappend_3510) => - (let stringappend_3520 = - (string_drop stringappend_3490 - stringappend_3510) in + Some (stringappend_3650,stringappend_3660) => + (let stringappend_3670 = + (string_drop stringappend_3640 stringappend_3660) in + if ((case ((sep_matches_prefix stringappend_3670)) of + Some (stringappend_3680,stringappend_3690) => + (let stringappend_3700 = + (string_drop stringappend_3670 + stringappend_3690) in if ((case ((reg_name_matches_prefix - stringappend_3520 + stringappend_3700 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3530,stringappend_3540) => - (case ((string_drop stringappend_3520 stringappend_3540)) of s0 => True ) + Some (stringappend_3710,stringappend_3720) => + (case ((string_drop stringappend_3700 stringappend_3720)) of s0 => True ) | None => False )) then True @@ -21414,46 +21751,46 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_3330 = (string_drop stringappend_00 ((string_length (''div'')))) in - (let (s, stringappend_3350) = - ((case ((maybe_not_u_matches_prefix stringappend_3330)) of - Some (stringappend_3340,stringappend_3350) => (stringappend_3340, stringappend_3350) + (let stringappend_3510 = (string_drop stringappend_00 ((string_length (''div'')))) in + (let (s, stringappend_3530) = + ((case ((maybe_not_u_matches_prefix stringappend_3510)) of + Some (stringappend_3520,stringappend_3530) => (stringappend_3520, stringappend_3530) )) in - (let stringappend_3360 = (string_drop stringappend_3330 stringappend_3350) in - (let stringappend_3370 = (string_drop stringappend_3360 ((string_length (''w'')))) in + (let stringappend_3540 = (string_drop stringappend_3510 stringappend_3530) in + (let stringappend_3550 = (string_drop stringappend_3540 ((string_length (''w'')))) in (case - (case ((spc_matches_prefix stringappend_3370)) of - Some (stringappend_3380,stringappend_3390) => (stringappend_3380, stringappend_3390) + (case ((spc_matches_prefix stringappend_3550)) of + Some (stringappend_3560,stringappend_3570) => (stringappend_3560, stringappend_3570) ) of - (_, stringappend_3390) => - (let stringappend_3400 = (string_drop stringappend_3370 stringappend_3390) in - (let (rd, stringappend_3420) = - ((case ((reg_name_matches_prefix stringappend_3400 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3410,stringappend_3420) => (stringappend_3410, stringappend_3420) + (_, stringappend_3570) => + (let stringappend_3580 = (string_drop stringappend_3550 stringappend_3570) in + (let (rd, stringappend_3600) = + ((case ((reg_name_matches_prefix stringappend_3580 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3590,stringappend_3600) => (stringappend_3590, stringappend_3600) )) in - (let stringappend_3430 = (string_drop stringappend_3400 stringappend_3420) in + (let stringappend_3610 = (string_drop stringappend_3580 stringappend_3600) in (case - (case ((sep_matches_prefix stringappend_3430)) of - Some (stringappend_3440,stringappend_3450) => (stringappend_3440, stringappend_3450) + (case ((sep_matches_prefix stringappend_3610)) of + Some (stringappend_3620,stringappend_3630) => (stringappend_3620, stringappend_3630) ) of - (_, stringappend_3450) => - (let stringappend_3460 = (string_drop stringappend_3430 stringappend_3450) in - (let (rs1, stringappend_3480) = - ((case ((reg_name_matches_prefix stringappend_3460 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3470,stringappend_3480) => (stringappend_3470, stringappend_3480) + (_, stringappend_3630) => + (let stringappend_3640 = (string_drop stringappend_3610 stringappend_3630) in + (let (rs1, stringappend_3660) = + ((case ((reg_name_matches_prefix stringappend_3640 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3650,stringappend_3660) => (stringappend_3650, stringappend_3660) )) in - (let stringappend_3490 = (string_drop stringappend_3460 stringappend_3480) in + (let stringappend_3670 = (string_drop stringappend_3640 stringappend_3660) in (case - (case ((sep_matches_prefix stringappend_3490)) of - Some (stringappend_3500,stringappend_3510) => (stringappend_3500, stringappend_3510) + (case ((sep_matches_prefix stringappend_3670)) of + Some (stringappend_3680,stringappend_3690) => (stringappend_3680, stringappend_3690) ) of - (_, stringappend_3510) => - (let stringappend_3520 = (string_drop stringappend_3490 stringappend_3510) in - (let (rs2, stringappend_3540) = - ((case ((reg_name_matches_prefix stringappend_3520 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3530,stringappend_3540) => (stringappend_3530, stringappend_3540) + (_, stringappend_3690) => + (let stringappend_3700 = (string_drop stringappend_3670 stringappend_3690) in + (let (rs2, stringappend_3720) = + ((case ((reg_name_matches_prefix stringappend_3700 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3710,stringappend_3720) => (stringappend_3710, stringappend_3720) )) in - (case ((string_drop stringappend_3520 stringappend_3540)) of + (case ((string_drop stringappend_3700 stringappend_3720)) of s1 => Some (DIVW (rs2,rs1,rd,s), ((string_length arg0)) - ((string_length s1))) @@ -21461,40 +21798,40 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) ))))) - else if (((((string_startswith stringappend_00 (''rem''))) \<and> ((let stringappend_3560 = (string_drop stringappend_00 ((string_length (''rem'')))) in - if ((case ((maybe_not_u_matches_prefix stringappend_3560)) of - Some (stringappend_3570,stringappend_3580) => - (let stringappend_3590 = (string_drop stringappend_3560 stringappend_3580) in - if (((((string_startswith stringappend_3590 (''w''))) \<and> ((let stringappend_3600 = - (string_drop stringappend_3590 ((string_length (''w'')))) in - if ((case ((spc_matches_prefix stringappend_3600)) of - Some (stringappend_3610,stringappend_3620) => - (let stringappend_3630 = - (string_drop stringappend_3600 stringappend_3620) in - if ((case ((reg_name_matches_prefix stringappend_3630 + else if (((((string_startswith stringappend_00 (''rem''))) \<and> ((let stringappend_3740 = (string_drop stringappend_00 ((string_length (''rem'')))) in + if ((case ((maybe_not_u_matches_prefix stringappend_3740)) of + Some (stringappend_3750,stringappend_3760) => + (let stringappend_3770 = (string_drop stringappend_3740 stringappend_3760) in + if (((((string_startswith stringappend_3770 (''w''))) \<and> ((let stringappend_3780 = + (string_drop stringappend_3770 ((string_length (''w'')))) in + if ((case ((spc_matches_prefix stringappend_3780)) of + Some (stringappend_3790,stringappend_3800) => + (let stringappend_3810 = + (string_drop stringappend_3780 stringappend_3800) in + if ((case ((reg_name_matches_prefix stringappend_3810 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3640,stringappend_3650) => - (let stringappend_3660 = - (string_drop stringappend_3630 stringappend_3650) in - if ((case ((sep_matches_prefix stringappend_3660)) of - Some (stringappend_3670,stringappend_3680) => - (let stringappend_3690 = - (string_drop stringappend_3660 stringappend_3680) in - if ((case ((reg_name_matches_prefix stringappend_3690 + Some (stringappend_3820,stringappend_3830) => + (let stringappend_3840 = + (string_drop stringappend_3810 stringappend_3830) in + if ((case ((sep_matches_prefix stringappend_3840)) of + Some (stringappend_3850,stringappend_3860) => + (let stringappend_3870 = + (string_drop stringappend_3840 stringappend_3860) in + if ((case ((reg_name_matches_prefix stringappend_3870 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3700,stringappend_3710) => - (let stringappend_3720 = - (string_drop stringappend_3690 stringappend_3710) in - if ((case ((sep_matches_prefix stringappend_3720)) of - Some (stringappend_3730,stringappend_3740) => - (let stringappend_3750 = - (string_drop stringappend_3720 - stringappend_3740) in + Some (stringappend_3880,stringappend_3890) => + (let stringappend_3900 = + (string_drop stringappend_3870 stringappend_3890) in + if ((case ((sep_matches_prefix stringappend_3900)) of + Some (stringappend_3910,stringappend_3920) => + (let stringappend_3930 = + (string_drop stringappend_3900 + stringappend_3920) in if ((case ((reg_name_matches_prefix - stringappend_3750 + stringappend_3930 :: (( 5 Word.word * ii))option)) of - Some (stringappend_3760,stringappend_3770) => - (case ((string_drop stringappend_3750 stringappend_3770)) of s0 => True ) + Some (stringappend_3940,stringappend_3950) => + (case ((string_drop stringappend_3930 stringappend_3950)) of s0 => True ) | None => False )) then True @@ -21525,46 +21862,46 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_3560 = (string_drop stringappend_00 ((string_length (''rem'')))) in - (let (s, stringappend_3580) = - ((case ((maybe_not_u_matches_prefix stringappend_3560)) of - Some (stringappend_3570,stringappend_3580) => (stringappend_3570, stringappend_3580) + (let stringappend_3740 = (string_drop stringappend_00 ((string_length (''rem'')))) in + (let (s, stringappend_3760) = + ((case ((maybe_not_u_matches_prefix stringappend_3740)) of + Some (stringappend_3750,stringappend_3760) => (stringappend_3750, stringappend_3760) )) in - (let stringappend_3590 = (string_drop stringappend_3560 stringappend_3580) in - (let stringappend_3600 = (string_drop stringappend_3590 ((string_length (''w'')))) in + (let stringappend_3770 = (string_drop stringappend_3740 stringappend_3760) in + (let stringappend_3780 = (string_drop stringappend_3770 ((string_length (''w'')))) in (case - (case ((spc_matches_prefix stringappend_3600)) of - Some (stringappend_3610,stringappend_3620) => (stringappend_3610, stringappend_3620) + (case ((spc_matches_prefix stringappend_3780)) of + Some (stringappend_3790,stringappend_3800) => (stringappend_3790, stringappend_3800) ) of - (_, stringappend_3620) => - (let stringappend_3630 = (string_drop stringappend_3600 stringappend_3620) in - (let (rd, stringappend_3650) = - ((case ((reg_name_matches_prefix stringappend_3630 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3640,stringappend_3650) => (stringappend_3640, stringappend_3650) + (_, stringappend_3800) => + (let stringappend_3810 = (string_drop stringappend_3780 stringappend_3800) in + (let (rd, stringappend_3830) = + ((case ((reg_name_matches_prefix stringappend_3810 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3820,stringappend_3830) => (stringappend_3820, stringappend_3830) )) in - (let stringappend_3660 = (string_drop stringappend_3630 stringappend_3650) in + (let stringappend_3840 = (string_drop stringappend_3810 stringappend_3830) in (case - (case ((sep_matches_prefix stringappend_3660)) of - Some (stringappend_3670,stringappend_3680) => (stringappend_3670, stringappend_3680) + (case ((sep_matches_prefix stringappend_3840)) of + Some (stringappend_3850,stringappend_3860) => (stringappend_3850, stringappend_3860) ) of - (_, stringappend_3680) => - (let stringappend_3690 = (string_drop stringappend_3660 stringappend_3680) in - (let (rs1, stringappend_3710) = - ((case ((reg_name_matches_prefix stringappend_3690 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3700,stringappend_3710) => (stringappend_3700, stringappend_3710) + (_, stringappend_3860) => + (let stringappend_3870 = (string_drop stringappend_3840 stringappend_3860) in + (let (rs1, stringappend_3890) = + ((case ((reg_name_matches_prefix stringappend_3870 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3880,stringappend_3890) => (stringappend_3880, stringappend_3890) )) in - (let stringappend_3720 = (string_drop stringappend_3690 stringappend_3710) in + (let stringappend_3900 = (string_drop stringappend_3870 stringappend_3890) in (case - (case ((sep_matches_prefix stringappend_3720)) of - Some (stringappend_3730,stringappend_3740) => (stringappend_3730, stringappend_3740) + (case ((sep_matches_prefix stringappend_3900)) of + Some (stringappend_3910,stringappend_3920) => (stringappend_3910, stringappend_3920) ) of - (_, stringappend_3740) => - (let stringappend_3750 = (string_drop stringappend_3720 stringappend_3740) in - (let (rs2, stringappend_3770) = - ((case ((reg_name_matches_prefix stringappend_3750 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_3760,stringappend_3770) => (stringappend_3760, stringappend_3770) + (_, stringappend_3920) => + (let stringappend_3930 = (string_drop stringappend_3900 stringappend_3920) in + (let (rs2, stringappend_3950) = + ((case ((reg_name_matches_prefix stringappend_3930 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_3940,stringappend_3950) => (stringappend_3940, stringappend_3950) )) in - (case ((string_drop stringappend_3750 stringappend_3770)) of + (case ((string_drop stringappend_3930 stringappend_3950)) of s1 => Some (REMW (rs2,rs1,rd,s), ((string_length arg0)) - ((string_length s1))) @@ -21572,22 +21909,22 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) ))))) - else if (((((string_startswith stringappend_00 (''fence''))) \<and> ((let stringappend_3790 = (string_drop stringappend_00 ((string_length (''fence'')))) in - if ((case ((spc_matches_prefix stringappend_3790)) of - Some (stringappend_3800,stringappend_3810) => - (let stringappend_3820 = (string_drop stringappend_3790 stringappend_3810) in - if ((case ((fence_bits_matches_prefix stringappend_3820 + else if (((((string_startswith stringappend_00 (''fence''))) \<and> ((let stringappend_3970 = (string_drop stringappend_00 ((string_length (''fence'')))) in + if ((case ((spc_matches_prefix stringappend_3970)) of + Some (stringappend_3980,stringappend_3990) => + (let stringappend_4000 = (string_drop stringappend_3970 stringappend_3990) in + if ((case ((fence_bits_matches_prefix stringappend_4000 :: (( 4 Word.word * ii))option)) of - Some (stringappend_3830,stringappend_3840) => - (let stringappend_3850 = (string_drop stringappend_3820 stringappend_3840) in - if ((case ((sep_matches_prefix stringappend_3850)) of - Some (stringappend_3860,stringappend_3870) => - (let stringappend_3880 = - (string_drop stringappend_3850 stringappend_3870) in - if ((case ((fence_bits_matches_prefix stringappend_3880 + Some (stringappend_4010,stringappend_4020) => + (let stringappend_4030 = (string_drop stringappend_4000 stringappend_4020) in + if ((case ((sep_matches_prefix stringappend_4030)) of + Some (stringappend_4040,stringappend_4050) => + (let stringappend_4060 = + (string_drop stringappend_4030 stringappend_4050) in + if ((case ((fence_bits_matches_prefix stringappend_4060 :: (( 4 Word.word * ii))option)) of - Some (stringappend_3890,stringappend_3900) => - (case ((string_drop stringappend_3880 stringappend_3900)) of s0 => True ) + Some (stringappend_4070,stringappend_4080) => + (case ((string_drop stringappend_4060 stringappend_4080)) of s0 => True ) | None => False )) then True @@ -21604,29 +21941,29 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_3790 = (string_drop stringappend_00 ((string_length (''fence'')))) in + (let stringappend_3970 = (string_drop stringappend_00 ((string_length (''fence'')))) in (case - (case ((spc_matches_prefix stringappend_3790)) of - Some (stringappend_3800,stringappend_3810) => (stringappend_3800, stringappend_3810) + (case ((spc_matches_prefix stringappend_3970)) of + Some (stringappend_3980,stringappend_3990) => (stringappend_3980, stringappend_3990) ) of - (_, stringappend_3810) => - (let stringappend_3820 = (string_drop stringappend_3790 stringappend_3810) in - (let (pred, stringappend_3840) = - ((case ((fence_bits_matches_prefix stringappend_3820 :: (( 4 Word.word * ii)) option)) of - Some (stringappend_3830,stringappend_3840) => (stringappend_3830, stringappend_3840) + (_, stringappend_3990) => + (let stringappend_4000 = (string_drop stringappend_3970 stringappend_3990) in + (let (pred, stringappend_4020) = + ((case ((fence_bits_matches_prefix stringappend_4000 :: (( 4 Word.word * ii)) option)) of + Some (stringappend_4010,stringappend_4020) => (stringappend_4010, stringappend_4020) )) in - (let stringappend_3850 = (string_drop stringappend_3820 stringappend_3840) in + (let stringappend_4030 = (string_drop stringappend_4000 stringappend_4020) in (case - (case ((sep_matches_prefix stringappend_3850)) of - Some (stringappend_3860,stringappend_3870) => (stringappend_3860, stringappend_3870) + (case ((sep_matches_prefix stringappend_4030)) of + Some (stringappend_4040,stringappend_4050) => (stringappend_4040, stringappend_4050) ) of - (_, stringappend_3870) => - (let stringappend_3880 = (string_drop stringappend_3850 stringappend_3870) in - (let (succ, stringappend_3900) = - ((case ((fence_bits_matches_prefix stringappend_3880 :: (( 4 Word.word * ii)) option)) of - Some (stringappend_3890,stringappend_3900) => (stringappend_3890, stringappend_3900) + (_, stringappend_4050) => + (let stringappend_4060 = (string_drop stringappend_4030 stringappend_4050) in + (let (succ, stringappend_4080) = + ((case ((fence_bits_matches_prefix stringappend_4060 :: (( 4 Word.word * ii)) option)) of + Some (stringappend_4070,stringappend_4080) => (stringappend_4070, stringappend_4080) )) in - (case ((string_drop stringappend_3880 stringappend_3900)) of + (case ((string_drop stringappend_4060 stringappend_4080)) of s0 => Some (FENCE (pred,succ), ((string_length arg0)) - ((string_length s0))) @@ -21675,22 +22012,22 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " (case ((string_drop stringappend_00 ((string_length (''wfi''))))) of s0 => Some (WFI () , ((string_length arg0)) - ((string_length s0))) ) - else if (((((string_startswith stringappend_00 (''sfence.vma''))) \<and> ((let stringappend_3980 = (string_drop stringappend_00 ((string_length (''sfence.vma'')))) in - if ((case ((spc_matches_prefix stringappend_3980)) of - Some (stringappend_3990,stringappend_4000) => - (let stringappend_4010 = (string_drop stringappend_3980 stringappend_4000) in - if ((case ((reg_name_matches_prefix stringappend_4010 + else if (((((string_startswith stringappend_00 (''sfence.vma''))) \<and> ((let stringappend_4160 = (string_drop stringappend_00 ((string_length (''sfence.vma'')))) in + if ((case ((spc_matches_prefix stringappend_4160)) of + Some (stringappend_4170,stringappend_4180) => + (let stringappend_4190 = (string_drop stringappend_4160 stringappend_4180) in + if ((case ((reg_name_matches_prefix stringappend_4190 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4020,stringappend_4030) => - (let stringappend_4040 = (string_drop stringappend_4010 stringappend_4030) in - if ((case ((sep_matches_prefix stringappend_4040)) of - Some (stringappend_4050,stringappend_4060) => - (let stringappend_4070 = - (string_drop stringappend_4040 stringappend_4060) in - if ((case ((reg_name_matches_prefix stringappend_4070 + Some (stringappend_4200,stringappend_4210) => + (let stringappend_4220 = (string_drop stringappend_4190 stringappend_4210) in + if ((case ((sep_matches_prefix stringappend_4220)) of + Some (stringappend_4230,stringappend_4240) => + (let stringappend_4250 = + (string_drop stringappend_4220 stringappend_4240) in + if ((case ((reg_name_matches_prefix stringappend_4250 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4080,stringappend_4090) => - (case ((string_drop stringappend_4070 stringappend_4090)) of s0 => True ) + Some (stringappend_4260,stringappend_4270) => + (case ((string_drop stringappend_4250 stringappend_4270)) of s0 => True ) | None => False )) then True @@ -21707,63 +22044,63 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_3980 = (string_drop stringappend_00 ((string_length (''sfence.vma'')))) in + (let stringappend_4160 = (string_drop stringappend_00 ((string_length (''sfence.vma'')))) in (case - (case ((spc_matches_prefix stringappend_3980)) of - Some (stringappend_3990,stringappend_4000) => (stringappend_3990, stringappend_4000) + (case ((spc_matches_prefix stringappend_4160)) of + Some (stringappend_4170,stringappend_4180) => (stringappend_4170, stringappend_4180) ) of - (_, stringappend_4000) => - (let stringappend_4010 = (string_drop stringappend_3980 stringappend_4000) in - (let (rs1, stringappend_4030) = - ((case ((reg_name_matches_prefix stringappend_4010 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4020,stringappend_4030) => (stringappend_4020, stringappend_4030) + (_, stringappend_4180) => + (let stringappend_4190 = (string_drop stringappend_4160 stringappend_4180) in + (let (rs1, stringappend_4210) = + ((case ((reg_name_matches_prefix stringappend_4190 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_4200,stringappend_4210) => (stringappend_4200, stringappend_4210) )) in - (let stringappend_4040 = (string_drop stringappend_4010 stringappend_4030) in + (let stringappend_4220 = (string_drop stringappend_4190 stringappend_4210) in (case - (case ((sep_matches_prefix stringappend_4040)) of - Some (stringappend_4050,stringappend_4060) => (stringappend_4050, stringappend_4060) + (case ((sep_matches_prefix stringappend_4220)) of + Some (stringappend_4230,stringappend_4240) => (stringappend_4230, stringappend_4240) ) of - (_, stringappend_4060) => - (let stringappend_4070 = (string_drop stringappend_4040 stringappend_4060) in - (let (rs2, stringappend_4090) = - ((case ((reg_name_matches_prefix stringappend_4070 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4080,stringappend_4090) => (stringappend_4080, stringappend_4090) + (_, stringappend_4240) => + (let stringappend_4250 = (string_drop stringappend_4220 stringappend_4240) in + (let (rs2, stringappend_4270) = + ((case ((reg_name_matches_prefix stringappend_4250 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_4260,stringappend_4270) => (stringappend_4260, stringappend_4270) )) in - (case ((string_drop stringappend_4070 stringappend_4090)) of + (case ((string_drop stringappend_4250 stringappend_4270)) of s0 => Some (SFENCE_VMA (rs1,rs2), ((string_length arg0)) - ((string_length s0))) ))) )))) )) - else if (((((string_startswith stringappend_00 (''lr.''))) \<and> ((let stringappend_4110 = (string_drop stringappend_00 ((string_length (''lr.'')))) in - if ((case ((maybe_aq_matches_prefix stringappend_4110)) of - Some (stringappend_4120,stringappend_4130) => - (let stringappend_4140 = (string_drop stringappend_4110 stringappend_4130) in - if ((case ((maybe_rl_matches_prefix stringappend_4140)) of - Some (stringappend_4150,stringappend_4160) => - (let stringappend_4170 = (string_drop stringappend_4140 stringappend_4160) in - if ((case ((size_mnemonic_matches_prefix stringappend_4170)) of - Some (stringappend_4180,stringappend_4190) => - (let stringappend_4200 = - (string_drop stringappend_4170 stringappend_4190) in - if ((case ((spc_matches_prefix stringappend_4200)) of - Some (stringappend_4210,stringappend_4220) => - (let stringappend_4230 = - (string_drop stringappend_4200 stringappend_4220) in - if ((case ((reg_name_matches_prefix stringappend_4230 + else if (((((string_startswith stringappend_00 (''lr.''))) \<and> ((let stringappend_4290 = (string_drop stringappend_00 ((string_length (''lr.'')))) in + if ((case ((maybe_aq_matches_prefix stringappend_4290)) of + Some (stringappend_4300,stringappend_4310) => + (let stringappend_4320 = (string_drop stringappend_4290 stringappend_4310) in + if ((case ((maybe_rl_matches_prefix stringappend_4320)) of + Some (stringappend_4330,stringappend_4340) => + (let stringappend_4350 = (string_drop stringappend_4320 stringappend_4340) in + if ((case ((size_mnemonic_matches_prefix stringappend_4350)) of + Some (stringappend_4360,stringappend_4370) => + (let stringappend_4380 = + (string_drop stringappend_4350 stringappend_4370) in + if ((case ((spc_matches_prefix stringappend_4380)) of + Some (stringappend_4390,stringappend_4400) => + (let stringappend_4410 = + (string_drop stringappend_4380 stringappend_4400) in + if ((case ((reg_name_matches_prefix stringappend_4410 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4240,stringappend_4250) => - (let stringappend_4260 = - (string_drop stringappend_4230 stringappend_4250) in - if ((case ((sep_matches_prefix stringappend_4260)) of - Some (stringappend_4270,stringappend_4280) => - (let stringappend_4290 = - (string_drop stringappend_4260 stringappend_4280) in - if ((case ((reg_name_matches_prefix stringappend_4290 + Some (stringappend_4420,stringappend_4430) => + (let stringappend_4440 = + (string_drop stringappend_4410 stringappend_4430) in + if ((case ((sep_matches_prefix stringappend_4440)) of + Some (stringappend_4450,stringappend_4460) => + (let stringappend_4470 = + (string_drop stringappend_4440 stringappend_4460) in + if ((case ((reg_name_matches_prefix stringappend_4470 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4300,stringappend_4310) => - (case ((string_drop stringappend_4290 stringappend_4310)) of s0 => True ) + Some (stringappend_4480,stringappend_4490) => + (case ((string_drop stringappend_4470 stringappend_4490)) of s0 => True ) | None => False )) then True @@ -21792,44 +22129,44 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_4110 = (string_drop stringappend_00 ((string_length (''lr.'')))) in - (let (aq, stringappend_4130) = - ((case ((maybe_aq_matches_prefix stringappend_4110)) of - Some (stringappend_4120,stringappend_4130) => (stringappend_4120, stringappend_4130) + (let stringappend_4290 = (string_drop stringappend_00 ((string_length (''lr.'')))) in + (let (aq, stringappend_4310) = + ((case ((maybe_aq_matches_prefix stringappend_4290)) of + Some (stringappend_4300,stringappend_4310) => (stringappend_4300, stringappend_4310) )) in - (let stringappend_4140 = (string_drop stringappend_4110 stringappend_4130) in - (let (rl, stringappend_4160) = - ((case ((maybe_rl_matches_prefix stringappend_4140)) of - Some (stringappend_4150,stringappend_4160) => (stringappend_4150, stringappend_4160) + (let stringappend_4320 = (string_drop stringappend_4290 stringappend_4310) in + (let (rl, stringappend_4340) = + ((case ((maybe_rl_matches_prefix stringappend_4320)) of + Some (stringappend_4330,stringappend_4340) => (stringappend_4330, stringappend_4340) )) in - (let stringappend_4170 = (string_drop stringappend_4140 stringappend_4160) in - (let (size1, stringappend_4190) = - ((case ((size_mnemonic_matches_prefix stringappend_4170)) of - Some (stringappend_4180,stringappend_4190) => (stringappend_4180, stringappend_4190) + (let stringappend_4350 = (string_drop stringappend_4320 stringappend_4340) in + (let (size1, stringappend_4370) = + ((case ((size_mnemonic_matches_prefix stringappend_4350)) of + Some (stringappend_4360,stringappend_4370) => (stringappend_4360, stringappend_4370) )) in - (let stringappend_4200 = (string_drop stringappend_4170 stringappend_4190) in + (let stringappend_4380 = (string_drop stringappend_4350 stringappend_4370) in (case - (case ((spc_matches_prefix stringappend_4200)) of - Some (stringappend_4210,stringappend_4220) => (stringappend_4210, stringappend_4220) + (case ((spc_matches_prefix stringappend_4380)) of + Some (stringappend_4390,stringappend_4400) => (stringappend_4390, stringappend_4400) ) of - (_, stringappend_4220) => - (let stringappend_4230 = (string_drop stringappend_4200 stringappend_4220) in - (let (rd, stringappend_4250) = - ((case ((reg_name_matches_prefix stringappend_4230 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4240,stringappend_4250) => (stringappend_4240, stringappend_4250) + (_, stringappend_4400) => + (let stringappend_4410 = (string_drop stringappend_4380 stringappend_4400) in + (let (rd, stringappend_4430) = + ((case ((reg_name_matches_prefix stringappend_4410 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_4420,stringappend_4430) => (stringappend_4420, stringappend_4430) )) in - (let stringappend_4260 = (string_drop stringappend_4230 stringappend_4250) in + (let stringappend_4440 = (string_drop stringappend_4410 stringappend_4430) in (case - (case ((sep_matches_prefix stringappend_4260)) of - Some (stringappend_4270,stringappend_4280) => (stringappend_4270, stringappend_4280) + (case ((sep_matches_prefix stringappend_4440)) of + Some (stringappend_4450,stringappend_4460) => (stringappend_4450, stringappend_4460) ) of - (_, stringappend_4280) => - (let stringappend_4290 = (string_drop stringappend_4260 stringappend_4280) in - (let (rs1, stringappend_4310) = - ((case ((reg_name_matches_prefix stringappend_4290 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4300,stringappend_4310) => (stringappend_4300, stringappend_4310) + (_, stringappend_4460) => + (let stringappend_4470 = (string_drop stringappend_4440 stringappend_4460) in + (let (rs1, stringappend_4490) = + ((case ((reg_name_matches_prefix stringappend_4470 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_4480,stringappend_4490) => (stringappend_4480, stringappend_4490) )) in - (case ((string_drop stringappend_4290 stringappend_4310)) of + (case ((string_drop stringappend_4470 stringappend_4490)) of s0 => Some (LOADRES (aq,rl,rs1,size1,rd), ((string_length arg0)) - @@ -21837,47 +22174,47 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " ))) )))) )))))))) - else if (((((string_startswith stringappend_00 (''sc.''))) \<and> ((let stringappend_4330 = (string_drop stringappend_00 ((string_length (''sc.'')))) in - if ((case ((maybe_aq_matches_prefix stringappend_4330)) of - Some (stringappend_4340,stringappend_4350) => - (let stringappend_4360 = (string_drop stringappend_4330 stringappend_4350) in - if ((case ((maybe_rl_matches_prefix stringappend_4360)) of - Some (stringappend_4370,stringappend_4380) => - (let stringappend_4390 = (string_drop stringappend_4360 stringappend_4380) in - if ((case ((size_mnemonic_matches_prefix stringappend_4390)) of - Some (stringappend_4400,stringappend_4410) => - (let stringappend_4420 = - (string_drop stringappend_4390 stringappend_4410) in - if ((case ((spc_matches_prefix stringappend_4420)) of - Some (stringappend_4430,stringappend_4440) => - (let stringappend_4450 = - (string_drop stringappend_4420 stringappend_4440) in - if ((case ((reg_name_matches_prefix stringappend_4450 + else if (((((string_startswith stringappend_00 (''sc.''))) \<and> ((let stringappend_4510 = (string_drop stringappend_00 ((string_length (''sc.'')))) in + if ((case ((maybe_aq_matches_prefix stringappend_4510)) of + Some (stringappend_4520,stringappend_4530) => + (let stringappend_4540 = (string_drop stringappend_4510 stringappend_4530) in + if ((case ((maybe_rl_matches_prefix stringappend_4540)) of + Some (stringappend_4550,stringappend_4560) => + (let stringappend_4570 = (string_drop stringappend_4540 stringappend_4560) in + if ((case ((size_mnemonic_matches_prefix stringappend_4570)) of + Some (stringappend_4580,stringappend_4590) => + (let stringappend_4600 = + (string_drop stringappend_4570 stringappend_4590) in + if ((case ((spc_matches_prefix stringappend_4600)) of + Some (stringappend_4610,stringappend_4620) => + (let stringappend_4630 = + (string_drop stringappend_4600 stringappend_4620) in + if ((case ((reg_name_matches_prefix stringappend_4630 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4460,stringappend_4470) => - (let stringappend_4480 = - (string_drop stringappend_4450 stringappend_4470) in - if ((case ((sep_matches_prefix stringappend_4480)) of - Some (stringappend_4490,stringappend_4500) => - (let stringappend_4510 = - (string_drop stringappend_4480 stringappend_4500) in - if ((case ((reg_name_matches_prefix stringappend_4510 + Some (stringappend_4640,stringappend_4650) => + (let stringappend_4660 = + (string_drop stringappend_4630 stringappend_4650) in + if ((case ((sep_matches_prefix stringappend_4660)) of + Some (stringappend_4670,stringappend_4680) => + (let stringappend_4690 = + (string_drop stringappend_4660 stringappend_4680) in + if ((case ((reg_name_matches_prefix stringappend_4690 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4520,stringappend_4530) => - (let stringappend_4540 = - (string_drop stringappend_4510 - stringappend_4530) in - if ((case ((sep_matches_prefix stringappend_4540)) of - Some (stringappend_4550,stringappend_4560) => - (let stringappend_4570 = - (string_drop stringappend_4540 - stringappend_4560) in + Some (stringappend_4700,stringappend_4710) => + (let stringappend_4720 = + (string_drop stringappend_4690 + stringappend_4710) in + if ((case ((sep_matches_prefix stringappend_4720)) of + Some (stringappend_4730,stringappend_4740) => + (let stringappend_4750 = + (string_drop stringappend_4720 + stringappend_4740) in if ((case ((reg_name_matches_prefix - stringappend_4570 + stringappend_4750 :: (( 5 Word.word * ii))option)) of Some - (stringappend_4580,stringappend_4590) => - (case ((string_drop stringappend_4570 stringappend_4590)) of s0 => True ) + (stringappend_4760,stringappend_4770) => + (case ((string_drop stringappend_4750 stringappend_4770)) of s0 => True ) | None => False )) then True @@ -21914,55 +22251,55 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_4330 = (string_drop stringappend_00 ((string_length (''sc.'')))) in - (let (aq, stringappend_4350) = - ((case ((maybe_aq_matches_prefix stringappend_4330)) of - Some (stringappend_4340,stringappend_4350) => (stringappend_4340, stringappend_4350) + (let stringappend_4510 = (string_drop stringappend_00 ((string_length (''sc.'')))) in + (let (aq, stringappend_4530) = + ((case ((maybe_aq_matches_prefix stringappend_4510)) of + Some (stringappend_4520,stringappend_4530) => (stringappend_4520, stringappend_4530) )) in - (let stringappend_4360 = (string_drop stringappend_4330 stringappend_4350) in - (let (rl, stringappend_4380) = - ((case ((maybe_rl_matches_prefix stringappend_4360)) of - Some (stringappend_4370,stringappend_4380) => (stringappend_4370, stringappend_4380) + (let stringappend_4540 = (string_drop stringappend_4510 stringappend_4530) in + (let (rl, stringappend_4560) = + ((case ((maybe_rl_matches_prefix stringappend_4540)) of + Some (stringappend_4550,stringappend_4560) => (stringappend_4550, stringappend_4560) )) in - (let stringappend_4390 = (string_drop stringappend_4360 stringappend_4380) in - (let (size1, stringappend_4410) = - ((case ((size_mnemonic_matches_prefix stringappend_4390)) of - Some (stringappend_4400,stringappend_4410) => (stringappend_4400, stringappend_4410) + (let stringappend_4570 = (string_drop stringappend_4540 stringappend_4560) in + (let (size1, stringappend_4590) = + ((case ((size_mnemonic_matches_prefix stringappend_4570)) of + Some (stringappend_4580,stringappend_4590) => (stringappend_4580, stringappend_4590) )) in - (let stringappend_4420 = (string_drop stringappend_4390 stringappend_4410) in + (let stringappend_4600 = (string_drop stringappend_4570 stringappend_4590) in (case - (case ((spc_matches_prefix stringappend_4420)) of - Some (stringappend_4430,stringappend_4440) => (stringappend_4430, stringappend_4440) + (case ((spc_matches_prefix stringappend_4600)) of + Some (stringappend_4610,stringappend_4620) => (stringappend_4610, stringappend_4620) ) of - (_, stringappend_4440) => - (let stringappend_4450 = (string_drop stringappend_4420 stringappend_4440) in - (let (rd, stringappend_4470) = - ((case ((reg_name_matches_prefix stringappend_4450 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4460,stringappend_4470) => (stringappend_4460, stringappend_4470) + (_, stringappend_4620) => + (let stringappend_4630 = (string_drop stringappend_4600 stringappend_4620) in + (let (rd, stringappend_4650) = + ((case ((reg_name_matches_prefix stringappend_4630 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_4640,stringappend_4650) => (stringappend_4640, stringappend_4650) )) in - (let stringappend_4480 = (string_drop stringappend_4450 stringappend_4470) in + (let stringappend_4660 = (string_drop stringappend_4630 stringappend_4650) in (case - (case ((sep_matches_prefix stringappend_4480)) of - Some (stringappend_4490,stringappend_4500) => (stringappend_4490, stringappend_4500) + (case ((sep_matches_prefix stringappend_4660)) of + Some (stringappend_4670,stringappend_4680) => (stringappend_4670, stringappend_4680) ) of - (_, stringappend_4500) => - (let stringappend_4510 = (string_drop stringappend_4480 stringappend_4500) in - (let (rs1, stringappend_4530) = - ((case ((reg_name_matches_prefix stringappend_4510 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4520,stringappend_4530) => (stringappend_4520, stringappend_4530) + (_, stringappend_4680) => + (let stringappend_4690 = (string_drop stringappend_4660 stringappend_4680) in + (let (rs1, stringappend_4710) = + ((case ((reg_name_matches_prefix stringappend_4690 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_4700,stringappend_4710) => (stringappend_4700, stringappend_4710) )) in - (let stringappend_4540 = (string_drop stringappend_4510 stringappend_4530) in + (let stringappend_4720 = (string_drop stringappend_4690 stringappend_4710) in (case - (case ((sep_matches_prefix stringappend_4540)) of - Some (stringappend_4550,stringappend_4560) => (stringappend_4550, stringappend_4560) + (case ((sep_matches_prefix stringappend_4720)) of + Some (stringappend_4730,stringappend_4740) => (stringappend_4730, stringappend_4740) ) of - (_, stringappend_4560) => - (let stringappend_4570 = (string_drop stringappend_4540 stringappend_4560) in - (let (rs2, stringappend_4590) = - ((case ((reg_name_matches_prefix stringappend_4570 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4580,stringappend_4590) => (stringappend_4580, stringappend_4590) + (_, stringappend_4740) => + (let stringappend_4750 = (string_drop stringappend_4720 stringappend_4740) in + (let (rs2, stringappend_4770) = + ((case ((reg_name_matches_prefix stringappend_4750 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_4760,stringappend_4770) => (stringappend_4760, stringappend_4770) )) in - (case ((string_drop stringappend_4570 stringappend_4590)) of + (case ((string_drop stringappend_4750 stringappend_4770)) of s0 => Some (STORECON (aq,rl,rs2,rs1,size1,rd), @@ -21972,49 +22309,49 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))))))) else if ((case ((amo_mnemonic_matches_prefix stringappend_00)) of - Some (stringappend_4610,stringappend_4620) => - (let stringappend_4630 = (string_drop stringappend_00 stringappend_4620) in - if (((((string_startswith stringappend_4630 (''.''))) \<and> ((let stringappend_4640 = (string_drop stringappend_4630 ((string_length (''.'')))) in - if ((case ((size_mnemonic_matches_prefix stringappend_4640)) of - Some (stringappend_4650,stringappend_4660) => - (let stringappend_4670 = (string_drop stringappend_4640 stringappend_4660) in - if ((case ((maybe_aq_matches_prefix stringappend_4670)) of - Some (stringappend_4680,stringappend_4690) => - (let stringappend_4700 = (string_drop stringappend_4670 stringappend_4690) in - if ((case ((maybe_rl_matches_prefix stringappend_4700)) of - Some (stringappend_4710,stringappend_4720) => - (let stringappend_4730 = - (string_drop stringappend_4700 stringappend_4720) in - if ((case ((spc_matches_prefix stringappend_4730)) of - Some (stringappend_4740,stringappend_4750) => - (let stringappend_4760 = - (string_drop stringappend_4730 stringappend_4750) in - if ((case ((reg_name_matches_prefix stringappend_4760 + Some (stringappend_4790,stringappend_4800) => + (let stringappend_4810 = (string_drop stringappend_00 stringappend_4800) in + if (((((string_startswith stringappend_4810 (''.''))) \<and> ((let stringappend_4820 = (string_drop stringappend_4810 ((string_length (''.'')))) in + if ((case ((size_mnemonic_matches_prefix stringappend_4820)) of + Some (stringappend_4830,stringappend_4840) => + (let stringappend_4850 = (string_drop stringappend_4820 stringappend_4840) in + if ((case ((maybe_aq_matches_prefix stringappend_4850)) of + Some (stringappend_4860,stringappend_4870) => + (let stringappend_4880 = (string_drop stringappend_4850 stringappend_4870) in + if ((case ((maybe_rl_matches_prefix stringappend_4880)) of + Some (stringappend_4890,stringappend_4900) => + (let stringappend_4910 = + (string_drop stringappend_4880 stringappend_4900) in + if ((case ((spc_matches_prefix stringappend_4910)) of + Some (stringappend_4920,stringappend_4930) => + (let stringappend_4940 = + (string_drop stringappend_4910 stringappend_4930) in + if ((case ((reg_name_matches_prefix stringappend_4940 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4770,stringappend_4780) => - (let stringappend_4790 = - (string_drop stringappend_4760 stringappend_4780) in - if ((case ((sep_matches_prefix stringappend_4790)) of - Some (stringappend_4800,stringappend_4810) => - (let stringappend_4820 = - (string_drop stringappend_4790 stringappend_4810) in - if ((case ((reg_name_matches_prefix stringappend_4820 + Some (stringappend_4950,stringappend_4960) => + (let stringappend_4970 = + (string_drop stringappend_4940 stringappend_4960) in + if ((case ((sep_matches_prefix stringappend_4970)) of + Some (stringappend_4980,stringappend_4990) => + (let stringappend_5000 = + (string_drop stringappend_4970 stringappend_4990) in + if ((case ((reg_name_matches_prefix stringappend_5000 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4830,stringappend_4840) => - (let stringappend_4850 = - (string_drop stringappend_4820 - stringappend_4840) in - if ((case ((sep_matches_prefix stringappend_4850)) of - Some (stringappend_4860,stringappend_4870) => - (let stringappend_4880 = - (string_drop stringappend_4850 - stringappend_4870) in + Some (stringappend_5010,stringappend_5020) => + (let stringappend_5030 = + (string_drop stringappend_5000 + stringappend_5020) in + if ((case ((sep_matches_prefix stringappend_5030)) of + Some (stringappend_5040,stringappend_5050) => + (let stringappend_5060 = + (string_drop stringappend_5030 + stringappend_5050) in if ((case ((reg_name_matches_prefix - stringappend_4880 + stringappend_5060 :: (( 5 Word.word * ii))option)) of Some - (stringappend_4890,stringappend_4900) => - (case ((string_drop stringappend_4880 stringappend_4900)) of s0 => True ) + (stringappend_5070,stringappend_5080) => + (case ((string_drop stringappend_5060 stringappend_5080)) of s0 => True ) | None => False )) then True @@ -22055,60 +22392,60 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " else False) | None => False )) then - (let (op1, stringappend_4620) = + (let (op1, stringappend_4800) = ((case ((amo_mnemonic_matches_prefix stringappend_00)) of - Some (stringappend_4610,stringappend_4620) => (stringappend_4610, stringappend_4620) + Some (stringappend_4790,stringappend_4800) => (stringappend_4790, stringappend_4800) )) in - (let stringappend_4630 = (string_drop stringappend_00 stringappend_4620) in - (let stringappend_4640 = (string_drop stringappend_4630 ((string_length (''.'')))) in - (let (width, stringappend_4660) = - ((case ((size_mnemonic_matches_prefix stringappend_4640)) of - Some (stringappend_4650,stringappend_4660) => (stringappend_4650, stringappend_4660) + (let stringappend_4810 = (string_drop stringappend_00 stringappend_4800) in + (let stringappend_4820 = (string_drop stringappend_4810 ((string_length (''.'')))) in + (let (width, stringappend_4840) = + ((case ((size_mnemonic_matches_prefix stringappend_4820)) of + Some (stringappend_4830,stringappend_4840) => (stringappend_4830, stringappend_4840) )) in - (let stringappend_4670 = (string_drop stringappend_4640 stringappend_4660) in - (let (aq, stringappend_4690) = - ((case ((maybe_aq_matches_prefix stringappend_4670)) of - Some (stringappend_4680,stringappend_4690) => (stringappend_4680, stringappend_4690) + (let stringappend_4850 = (string_drop stringappend_4820 stringappend_4840) in + (let (aq, stringappend_4870) = + ((case ((maybe_aq_matches_prefix stringappend_4850)) of + Some (stringappend_4860,stringappend_4870) => (stringappend_4860, stringappend_4870) )) in - (let stringappend_4700 = (string_drop stringappend_4670 stringappend_4690) in - (let (rl, stringappend_4720) = - ((case ((maybe_rl_matches_prefix stringappend_4700)) of - Some (stringappend_4710,stringappend_4720) => (stringappend_4710, stringappend_4720) + (let stringappend_4880 = (string_drop stringappend_4850 stringappend_4870) in + (let (rl, stringappend_4900) = + ((case ((maybe_rl_matches_prefix stringappend_4880)) of + Some (stringappend_4890,stringappend_4900) => (stringappend_4890, stringappend_4900) )) in - (let stringappend_4730 = (string_drop stringappend_4700 stringappend_4720) in + (let stringappend_4910 = (string_drop stringappend_4880 stringappend_4900) in (case - (case ((spc_matches_prefix stringappend_4730)) of - Some (stringappend_4740,stringappend_4750) => (stringappend_4740, stringappend_4750) + (case ((spc_matches_prefix stringappend_4910)) of + Some (stringappend_4920,stringappend_4930) => (stringappend_4920, stringappend_4930) ) of - (_, stringappend_4750) => - (let stringappend_4760 = (string_drop stringappend_4730 stringappend_4750) in - (let (rd, stringappend_4780) = - ((case ((reg_name_matches_prefix stringappend_4760 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4770,stringappend_4780) => (stringappend_4770, stringappend_4780) + (_, stringappend_4930) => + (let stringappend_4940 = (string_drop stringappend_4910 stringappend_4930) in + (let (rd, stringappend_4960) = + ((case ((reg_name_matches_prefix stringappend_4940 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_4950,stringappend_4960) => (stringappend_4950, stringappend_4960) )) in - (let stringappend_4790 = (string_drop stringappend_4760 stringappend_4780) in + (let stringappend_4970 = (string_drop stringappend_4940 stringappend_4960) in (case - (case ((sep_matches_prefix stringappend_4790)) of - Some (stringappend_4800,stringappend_4810) => (stringappend_4800, stringappend_4810) + (case ((sep_matches_prefix stringappend_4970)) of + Some (stringappend_4980,stringappend_4990) => (stringappend_4980, stringappend_4990) ) of - (_, stringappend_4810) => - (let stringappend_4820 = (string_drop stringappend_4790 stringappend_4810) in - (let (rs1, stringappend_4840) = - ((case ((reg_name_matches_prefix stringappend_4820 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4830,stringappend_4840) => (stringappend_4830, stringappend_4840) + (_, stringappend_4990) => + (let stringappend_5000 = (string_drop stringappend_4970 stringappend_4990) in + (let (rs1, stringappend_5020) = + ((case ((reg_name_matches_prefix stringappend_5000 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5010,stringappend_5020) => (stringappend_5010, stringappend_5020) )) in - (let stringappend_4850 = (string_drop stringappend_4820 stringappend_4840) in + (let stringappend_5030 = (string_drop stringappend_5000 stringappend_5020) in (case - (case ((sep_matches_prefix stringappend_4850)) of - Some (stringappend_4860,stringappend_4870) => (stringappend_4860, stringappend_4870) + (case ((sep_matches_prefix stringappend_5030)) of + Some (stringappend_5040,stringappend_5050) => (stringappend_5040, stringappend_5050) ) of - (_, stringappend_4870) => - (let stringappend_4880 = (string_drop stringappend_4850 stringappend_4870) in - (let (rs2, stringappend_4900) = - ((case ((reg_name_matches_prefix stringappend_4880 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4890,stringappend_4900) => (stringappend_4890, stringappend_4900) + (_, stringappend_5050) => + (let stringappend_5060 = (string_drop stringappend_5030 stringappend_5050) in + (let (rs2, stringappend_5080) = + ((case ((reg_name_matches_prefix stringappend_5060 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5070,stringappend_5080) => (stringappend_5070, stringappend_5080) )) in - (case ((string_drop stringappend_4880 stringappend_4900)) of + (case ((string_drop stringappend_5060 stringappend_5080)) of s0 => Some (AMO (op1,aq,rl,rs2,rs1,width,rd), @@ -22118,34 +22455,34 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))))))))) else if ((case ((csr_mnemonic_matches_prefix stringappend_00)) of - Some (stringappend_4920,stringappend_4930) => - (let stringappend_4940 = (string_drop stringappend_00 stringappend_4930) in - if (((((string_startswith stringappend_4940 (''i''))) \<and> ((let stringappend_4950 = (string_drop stringappend_4940 ((string_length (''i'')))) in - if ((case ((spc_matches_prefix stringappend_4950)) of - Some (stringappend_4960,stringappend_4970) => - (let stringappend_4980 = (string_drop stringappend_4950 stringappend_4970) in - if ((case ((reg_name_matches_prefix stringappend_4980 + Some (stringappend_5100,stringappend_5110) => + (let stringappend_5120 = (string_drop stringappend_00 stringappend_5110) in + if (((((string_startswith stringappend_5120 (''i''))) \<and> ((let stringappend_5130 = (string_drop stringappend_5120 ((string_length (''i'')))) in + if ((case ((spc_matches_prefix stringappend_5130)) of + Some (stringappend_5140,stringappend_5150) => + (let stringappend_5160 = (string_drop stringappend_5130 stringappend_5150) in + if ((case ((reg_name_matches_prefix stringappend_5160 :: (( 5 Word.word * ii))option)) of - Some (stringappend_4990,stringappend_5000) => - (let stringappend_5010 = (string_drop stringappend_4980 stringappend_5000) in - if ((case ((sep_matches_prefix stringappend_5010)) of - Some (stringappend_5020,stringappend_5030) => - (let stringappend_5040 = - (string_drop stringappend_5010 stringappend_5030) in + Some (stringappend_5170,stringappend_5180) => + (let stringappend_5190 = (string_drop stringappend_5160 stringappend_5180) in + if ((case ((sep_matches_prefix stringappend_5190)) of + Some (stringappend_5200,stringappend_5210) => + (let stringappend_5220 = + (string_drop stringappend_5190 stringappend_5210) in if ((case ((hex_bits_5_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5040 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5220 :: (( 5 Word.word * ii))option)) of - Some (stringappend_5050,stringappend_5060) => - (let stringappend_5070 = - (string_drop stringappend_5040 stringappend_5060) in - if ((case ((sep_matches_prefix stringappend_5070)) of - Some (stringappend_5080,stringappend_5090) => - (let stringappend_5100 = - (string_drop stringappend_5070 stringappend_5090) in - if ((case ((csr_name_map_matches_prefix stringappend_5100 + Some (stringappend_5230,stringappend_5240) => + (let stringappend_5250 = + (string_drop stringappend_5220 stringappend_5240) in + if ((case ((sep_matches_prefix stringappend_5250)) of + Some (stringappend_5260,stringappend_5270) => + (let stringappend_5280 = + (string_drop stringappend_5250 stringappend_5270) in + if ((case ((csr_name_map_matches_prefix stringappend_5280 :: (( 12 Word.word * ii))option)) of - Some (stringappend_5110,stringappend_5120) => - (case ((string_drop stringappend_5100 stringappend_5120)) of s0 => True ) + Some (stringappend_5290,stringappend_5300) => + (case ((string_drop stringappend_5280 stringappend_5300)) of s0 => True ) | None => False )) then True @@ -22174,47 +22511,47 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " else False) | None => False )) then - (let (op1, stringappend_4930) = + (let (op1, stringappend_5110) = ((case ((csr_mnemonic_matches_prefix stringappend_00)) of - Some (stringappend_4920,stringappend_4930) => (stringappend_4920, stringappend_4930) + Some (stringappend_5100,stringappend_5110) => (stringappend_5100, stringappend_5110) )) in - (let stringappend_4940 = (string_drop stringappend_00 stringappend_4930) in - (let stringappend_4950 = (string_drop stringappend_4940 ((string_length (''i'')))) in + (let stringappend_5120 = (string_drop stringappend_00 stringappend_5110) in + (let stringappend_5130 = (string_drop stringappend_5120 ((string_length (''i'')))) in (case - (case ((spc_matches_prefix stringappend_4950)) of - Some (stringappend_4960,stringappend_4970) => (stringappend_4960, stringappend_4970) + (case ((spc_matches_prefix stringappend_5130)) of + Some (stringappend_5140,stringappend_5150) => (stringappend_5140, stringappend_5150) ) of - (_, stringappend_4970) => - (let stringappend_4980 = (string_drop stringappend_4950 stringappend_4970) in - (let (rd, stringappend_5000) = - ((case ((reg_name_matches_prefix stringappend_4980 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_4990,stringappend_5000) => (stringappend_4990, stringappend_5000) + (_, stringappend_5150) => + (let stringappend_5160 = (string_drop stringappend_5130 stringappend_5150) in + (let (rd, stringappend_5180) = + ((case ((reg_name_matches_prefix stringappend_5160 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5170,stringappend_5180) => (stringappend_5170, stringappend_5180) )) in - (let stringappend_5010 = (string_drop stringappend_4980 stringappend_5000) in + (let stringappend_5190 = (string_drop stringappend_5160 stringappend_5180) in (case - (case ((sep_matches_prefix stringappend_5010)) of - Some (stringappend_5020,stringappend_5030) => (stringappend_5020, stringappend_5030) + (case ((sep_matches_prefix stringappend_5190)) of + Some (stringappend_5200,stringappend_5210) => (stringappend_5200, stringappend_5210) ) of - (_, stringappend_5030) => - (let stringappend_5040 = (string_drop stringappend_5010 stringappend_5030) in - (let (rs1, stringappend_5060) = + (_, stringappend_5210) => + (let stringappend_5220 = (string_drop stringappend_5190 stringappend_5210) in + (let (rs1, stringappend_5240) = ((case ((hex_bits_5_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_5040 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_5050,stringappend_5060) => (stringappend_5050, stringappend_5060) + stringappend_5220 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5230,stringappend_5240) => (stringappend_5230, stringappend_5240) )) in - (let stringappend_5070 = (string_drop stringappend_5040 stringappend_5060) in + (let stringappend_5250 = (string_drop stringappend_5220 stringappend_5240) in (case - (case ((sep_matches_prefix stringappend_5070)) of - Some (stringappend_5080,stringappend_5090) => (stringappend_5080, stringappend_5090) + (case ((sep_matches_prefix stringappend_5250)) of + Some (stringappend_5260,stringappend_5270) => (stringappend_5260, stringappend_5270) ) of - (_, stringappend_5090) => - (let stringappend_5100 = (string_drop stringappend_5070 stringappend_5090) in - (let (csr, stringappend_5120) = - ((case ((csr_name_map_matches_prefix stringappend_5100 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_5110,stringappend_5120) => (stringappend_5110, stringappend_5120) + (_, stringappend_5270) => + (let stringappend_5280 = (string_drop stringappend_5250 stringappend_5270) in + (let (csr, stringappend_5300) = + ((case ((csr_name_map_matches_prefix stringappend_5280 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_5290,stringappend_5300) => (stringappend_5290, stringappend_5300) )) in - (case ((string_drop stringappend_5100 stringappend_5120)) of + (case ((string_drop stringappend_5280 stringappend_5300)) of s0 => Some (CSR (csr,rs1,rd,True,op1), ((string_length arg0)) - ((string_length s0))) @@ -22223,30 +22560,30 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) else if ((case ((csr_mnemonic_matches_prefix stringappend_00)) of - Some (stringappend_5140,stringappend_5150) => - (let stringappend_5160 = (string_drop stringappend_00 stringappend_5150) in - if ((case ((spc_matches_prefix stringappend_5160)) of - Some (stringappend_5170,stringappend_5180) => - (let stringappend_5190 = (string_drop stringappend_5160 stringappend_5180) in - if ((case ((reg_name_matches_prefix stringappend_5190 :: (( 5 Word.word * ii))option)) of - Some (stringappend_5200,stringappend_5210) => - (let stringappend_5220 = (string_drop stringappend_5190 stringappend_5210) in - if ((case ((sep_matches_prefix stringappend_5220)) of - Some (stringappend_5230,stringappend_5240) => - (let stringappend_5250 = (string_drop stringappend_5220 stringappend_5240) in - if ((case ((reg_name_matches_prefix stringappend_5250 + Some (stringappend_5320,stringappend_5330) => + (let stringappend_5340 = (string_drop stringappend_00 stringappend_5330) in + if ((case ((spc_matches_prefix stringappend_5340)) of + Some (stringappend_5350,stringappend_5360) => + (let stringappend_5370 = (string_drop stringappend_5340 stringappend_5360) in + if ((case ((reg_name_matches_prefix stringappend_5370 :: (( 5 Word.word * ii))option)) of + Some (stringappend_5380,stringappend_5390) => + (let stringappend_5400 = (string_drop stringappend_5370 stringappend_5390) in + if ((case ((sep_matches_prefix stringappend_5400)) of + Some (stringappend_5410,stringappend_5420) => + (let stringappend_5430 = (string_drop stringappend_5400 stringappend_5420) in + if ((case ((reg_name_matches_prefix stringappend_5430 :: (( 5 Word.word * ii))option)) of - Some (stringappend_5260,stringappend_5270) => - (let stringappend_5280 = - (string_drop stringappend_5250 stringappend_5270) in - if ((case ((sep_matches_prefix stringappend_5280)) of - Some (stringappend_5290,stringappend_5300) => - (let stringappend_5310 = - (string_drop stringappend_5280 stringappend_5300) in - if ((case ((csr_name_map_matches_prefix stringappend_5310 + Some (stringappend_5440,stringappend_5450) => + (let stringappend_5460 = + (string_drop stringappend_5430 stringappend_5450) in + if ((case ((sep_matches_prefix stringappend_5460)) of + Some (stringappend_5470,stringappend_5480) => + (let stringappend_5490 = + (string_drop stringappend_5460 stringappend_5480) in + if ((case ((csr_name_map_matches_prefix stringappend_5490 :: (( 12 Word.word * ii))option)) of - Some (stringappend_5320,stringappend_5330) => - (case ((string_drop stringappend_5310 stringappend_5330)) of s0 => True ) + Some (stringappend_5500,stringappend_5510) => + (case ((string_drop stringappend_5490 stringappend_5510)) of s0 => True ) | None => False )) then True @@ -22273,44 +22610,44 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " else False) | None => False )) then - (let (op1, stringappend_5150) = + (let (op1, stringappend_5330) = ((case ((csr_mnemonic_matches_prefix stringappend_00)) of - Some (stringappend_5140,stringappend_5150) => (stringappend_5140, stringappend_5150) + Some (stringappend_5320,stringappend_5330) => (stringappend_5320, stringappend_5330) )) in - (let stringappend_5160 = (string_drop stringappend_00 stringappend_5150) in + (let stringappend_5340 = (string_drop stringappend_00 stringappend_5330) in (case - (case ((spc_matches_prefix stringappend_5160)) of - Some (stringappend_5170,stringappend_5180) => (stringappend_5170, stringappend_5180) + (case ((spc_matches_prefix stringappend_5340)) of + Some (stringappend_5350,stringappend_5360) => (stringappend_5350, stringappend_5360) ) of - (_, stringappend_5180) => - (let stringappend_5190 = (string_drop stringappend_5160 stringappend_5180) in - (let (rd, stringappend_5210) = - ((case ((reg_name_matches_prefix stringappend_5190 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_5200,stringappend_5210) => (stringappend_5200, stringappend_5210) + (_, stringappend_5360) => + (let stringappend_5370 = (string_drop stringappend_5340 stringappend_5360) in + (let (rd, stringappend_5390) = + ((case ((reg_name_matches_prefix stringappend_5370 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5380,stringappend_5390) => (stringappend_5380, stringappend_5390) )) in - (let stringappend_5220 = (string_drop stringappend_5190 stringappend_5210) in + (let stringappend_5400 = (string_drop stringappend_5370 stringappend_5390) in (case - (case ((sep_matches_prefix stringappend_5220)) of - Some (stringappend_5230,stringappend_5240) => (stringappend_5230, stringappend_5240) + (case ((sep_matches_prefix stringappend_5400)) of + Some (stringappend_5410,stringappend_5420) => (stringappend_5410, stringappend_5420) ) of - (_, stringappend_5240) => - (let stringappend_5250 = (string_drop stringappend_5220 stringappend_5240) in - (let (rs1, stringappend_5270) = - ((case ((reg_name_matches_prefix stringappend_5250 :: (( 5 Word.word * ii)) option)) of - Some (stringappend_5260,stringappend_5270) => (stringappend_5260, stringappend_5270) + (_, stringappend_5420) => + (let stringappend_5430 = (string_drop stringappend_5400 stringappend_5420) in + (let (rs1, stringappend_5450) = + ((case ((reg_name_matches_prefix stringappend_5430 :: (( 5 Word.word * ii)) option)) of + Some (stringappend_5440,stringappend_5450) => (stringappend_5440, stringappend_5450) )) in - (let stringappend_5280 = (string_drop stringappend_5250 stringappend_5270) in + (let stringappend_5460 = (string_drop stringappend_5430 stringappend_5450) in (case - (case ((sep_matches_prefix stringappend_5280)) of - Some (stringappend_5290,stringappend_5300) => (stringappend_5290, stringappend_5300) + (case ((sep_matches_prefix stringappend_5460)) of + Some (stringappend_5470,stringappend_5480) => (stringappend_5470, stringappend_5480) ) of - (_, stringappend_5300) => - (let stringappend_5310 = (string_drop stringappend_5280 stringappend_5300) in - (let (csr, stringappend_5330) = - ((case ((csr_name_map_matches_prefix stringappend_5310 :: (( 12 Word.word * ii)) option)) of - Some (stringappend_5320,stringappend_5330) => (stringappend_5320, stringappend_5330) + (_, stringappend_5480) => + (let stringappend_5490 = (string_drop stringappend_5460 stringappend_5480) in + (let (csr, stringappend_5510) = + ((case ((csr_name_map_matches_prefix stringappend_5490 :: (( 12 Word.word * ii)) option)) of + Some (stringappend_5500,stringappend_5510) => (stringappend_5500, stringappend_5510) )) in - (case ((string_drop stringappend_5310 stringappend_5330)) of + (case ((string_drop stringappend_5490 stringappend_5510)) of s0 => Some (CSR (csr,rs1,rd,False,op1), ((string_length arg0)) - @@ -22319,15 +22656,15 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )))) )))) ))) - else if (((((string_startswith stringappend_00 (''illegal''))) \<and> ((let stringappend_5350 = (string_drop stringappend_00 ((string_length (''illegal'')))) in - if ((case ((spc_matches_prefix stringappend_5350)) of - Some (stringappend_5360,stringappend_5370) => - (let stringappend_5380 = (string_drop stringappend_5350 stringappend_5370) in + else if (((((string_startswith stringappend_00 (''illegal''))) \<and> ((let stringappend_5530 = (string_drop stringappend_00 ((string_length (''illegal'')))) in + if ((case ((spc_matches_prefix stringappend_5530)) of + Some (stringappend_5540,stringappend_5550) => + (let stringappend_5560 = (string_drop stringappend_5530 stringappend_5550) in if ((case ((hex_bits_32_matches_prefix - instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5380 + instance_Sail2_values_Bitvector_Machine_word_mword_dict stringappend_5560 :: (( 32 Word.word * ii))option)) of - Some (stringappend_5390,stringappend_5400) => - (case ((string_drop stringappend_5380 stringappend_5400)) of s0 => True ) + Some (stringappend_5570,stringappend_5580) => + (case ((string_drop stringappend_5560 stringappend_5580)) of s0 => True ) | None => False )) then True @@ -22336,20 +22673,20 @@ definition assembly_matches_prefix :: " string \<Rightarrow>(ast*int)option " )) then True else False))))) then - (let stringappend_5350 = (string_drop stringappend_00 ((string_length (''illegal'')))) in + (let stringappend_5530 = (string_drop stringappend_00 ((string_length (''illegal'')))) in (case - (case ((spc_matches_prefix stringappend_5350)) of - Some (stringappend_5360,stringappend_5370) => (stringappend_5360, stringappend_5370) + (case ((spc_matches_prefix stringappend_5530)) of + Some (stringappend_5540,stringappend_5550) => (stringappend_5540, stringappend_5550) ) of - (_, stringappend_5370) => - (let stringappend_5380 = (string_drop stringappend_5350 stringappend_5370) in - (let (s, stringappend_5400) = + (_, stringappend_5550) => + (let stringappend_5560 = (string_drop stringappend_5530 stringappend_5550) in + (let (s, stringappend_5580) = ((case ((hex_bits_32_matches_prefix instance_Sail2_values_Bitvector_Machine_word_mword_dict - stringappend_5380 :: (( 32 Word.word * ii)) option)) of - Some (stringappend_5390,stringappend_5400) => (stringappend_5390, stringappend_5400) + stringappend_5560 :: (( 32 Word.word * ii)) option)) of + Some (stringappend_5570,stringappend_5580) => (stringappend_5570, stringappend_5580) )) in - (case ((string_drop stringappend_5380 stringappend_5400)) of + (case ((string_drop stringappend_5560 stringappend_5580)) of s1 => Some (ILLEGAL s, ((string_length arg0)) - ((string_length s1))) ))) )) @@ -22364,8 +22701,9 @@ fun encdec_forwards :: " ast \<Rightarrow>(32)Word.word " where :: 32 Word.word))" |" encdec_forwards (RISCV_JAL (v__172,rd)) = ( (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__172 (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word)) in - (let (imm_7_0 :: 8 bits) = ((subrange_vec_dec v__172 (( 19 :: int)::ii) (( 12 :: int)::ii) :: 8 Word.word)) in (let (imm_8 :: 1 bits) = ((subrange_vec_dec v__172 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in + (let (imm_7_0 :: 8 bits) = ((subrange_vec_dec v__172 (( 19 :: int)::ii) (( 12 :: int)::ii) :: 8 Word.word)) in + (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__172 (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word)) in (let (imm_18_13 :: 6 bits) = ((subrange_vec_dec v__172 (( 10 :: int)::ii) (( 5 :: int)::ii) :: 6 Word.word)) in (let (imm_12_9 :: 4 bits) = ((subrange_vec_dec v__172 (( 4 :: int)::ii) (( 1 :: int)::ii) :: 4 Word.word)) in (concat_vec imm_19 @@ -22379,7 +22717,7 @@ fun encdec_forwards :: " ast \<Rightarrow>(32)Word.word " where :: 21 Word.word)) :: 25 Word.word)) :: 31 Word.word)) - :: 32 Word.word)))))))" + :: 32 Word.word))))))))" |" encdec_forwards (RISCV_JALR (imm,rs1,rd)) = ( (concat_vec imm ((concat_vec rs1 @@ -22388,11 +22726,12 @@ fun encdec_forwards :: " ast \<Rightarrow>(32)Word.word " where :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word))" -|" encdec_forwards (BTYPE (v__173,rs2,rs1,op1)) = ( - (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__173 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (imm5_0 :: 1 bits) = ((subrange_vec_dec v__173 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in - (let (imm7_5_0 :: 6 bits) = ((subrange_vec_dec v__173 (( 10 :: int)::ii) (( 5 :: int)::ii) :: 6 Word.word)) in - (let (imm5_4_1 :: 4 bits) = ((subrange_vec_dec v__173 (( 4 :: int)::ii) (( 1 :: int)::ii) :: 4 Word.word)) in +|" encdec_forwards (BTYPE (v__174,rs2,rs1,op1)) = ( + (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__174 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__174 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (let (imm7_5_0 :: 6 bits) = ((subrange_vec_dec v__174 (( 10 :: int)::ii) (( 5 :: int)::ii) :: 6 Word.word)) in + (let (imm5_4_1 :: 4 bits) = ((subrange_vec_dec v__174 (( 4 :: int)::ii) (( 1 :: int)::ii) :: 4 Word.word)) in + (let (imm5_0 :: 1 bits) = ((subrange_vec_dec v__174 (( 11 :: int)::ii) (( 11 :: int)::ii) :: 1 Word.word)) in (concat_vec imm7_6 ((concat_vec imm7_5_0 ((concat_vec rs2 @@ -22406,7 +22745,7 @@ fun encdec_forwards :: " ast \<Rightarrow>(32)Word.word " where :: 20 Word.word)) :: 25 Word.word)) :: 31 Word.word)) - :: 32 Word.word))))))" + :: 32 Word.word)))))))" |" encdec_forwards (ITYPE (imm,rs1,rd,op1)) = ( (concat_vec imm ((concat_vec rs1 @@ -22555,9 +22894,10 @@ fun encdec_forwards :: " ast \<Rightarrow>(32)Word.word " where :: 15 Word.word)) :: 20 Word.word)) :: 32 Word.word))" -|" encdec_forwards (STORE (v__174,rs2,rs1,size1,False,False)) = ( - (let (imm7 :: 7 bits) = ((subrange_vec_dec v__174 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in - (let (imm5 :: 5 bits) = ((subrange_vec_dec v__174 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in +|" encdec_forwards (STORE (v__176,rs2,rs1,size1,False,False)) = ( + (let (imm7 :: 7 bits) = ((subrange_vec_dec v__176 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in + (let (imm7 :: 7 bits) = ((subrange_vec_dec v__176 (( 11 :: int)::ii) (( 5 :: int)::ii) :: 7 Word.word)) in + (let (imm5 :: 5 bits) = ((subrange_vec_dec v__176 (( 4 :: int)::ii) (( 0 :: int)::ii) :: 5 Word.word)) in (concat_vec imm7 ((concat_vec rs2 ((concat_vec rs1 @@ -22569,7 +22909,7 @@ fun encdec_forwards :: " ast \<Rightarrow>(32)Word.word " where :: 15 Word.word)) :: 20 Word.word)) :: 25 Word.word)) - :: 32 Word.word))))" + :: 32 Word.word)))))" |" encdec_forwards (ADDIW (imm,rs1,rd)) = ( (concat_vec imm ((concat_vec rs1 @@ -22658,6 +22998,36 @@ fun encdec_forwards :: " ast \<Rightarrow>(32)Word.word " where :: 20 Word.word)) :: 25 Word.word)) :: 32 Word.word))" +|" encdec_forwards (SHIFTIWOP (shamt,rs1,rd,RISCV_SLLIW)) = ( + (concat_vec (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word) + ((concat_vec shamt + ((concat_vec rs1 + ((concat_vec (vec_of_bits [B0,B0,B1] :: 3 Word.word) + ((concat_vec rd (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word) :: 12 Word.word)) + :: 15 Word.word)) + :: 20 Word.word)) + :: 25 Word.word)) + :: 32 Word.word))" +|" encdec_forwards (SHIFTIWOP (shamt,rs1,rd,RISCV_SRLIW)) = ( + (concat_vec (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word) + ((concat_vec shamt + ((concat_vec rs1 + ((concat_vec (vec_of_bits [B1,B0,B1] :: 3 Word.word) + ((concat_vec rd (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word) :: 12 Word.word)) + :: 15 Word.word)) + :: 20 Word.word)) + :: 25 Word.word)) + :: 32 Word.word))" +|" encdec_forwards (SHIFTIWOP (shamt,rs1,rd,RISCV_SRAIW)) = ( + (concat_vec (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word) + ((concat_vec shamt + ((concat_vec rs1 + ((concat_vec (vec_of_bits [B1,B0,B1] :: 3 Word.word) + ((concat_vec rd (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word) :: 12 Word.word)) + :: 15 Word.word)) + :: 20 Word.word)) + :: 25 Word.word)) + :: 32 Word.word))" |" encdec_forwards (MUL (rs2,rs1,rd,high,signed1,signed2)) = ( (concat_vec (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word) ((concat_vec rs2 @@ -22912,22 +23282,26 @@ fun encdec_forwards :: " ast \<Rightarrow>(32)Word.word " where definition encdec_backwards :: "(32)Word.word \<Rightarrow> ast " where " encdec_backwards arg0 = ( - (let v__175 = arg0 in - if ((let mappingpatterns_230 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - encdec_uop_backwards_matches mappingpatterns_230)) then - (let (imm :: 20 Word.word) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 12 :: int)::ii) :: 20 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_230 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in + (let v__177 = arg0 in + if ((let mappingpatterns_230 = ((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in + (((encdec_uop_backwards_matches mappingpatterns_230)) \<and> (if ((encdec_uop_backwards_matches mappingpatterns_230)) then + (let op1 = (encdec_uop_backwards mappingpatterns_230) in + True) + else False)))) then + (let (imm :: 20 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 12 :: int)::ii) :: 20 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm :: 20 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 12 :: int)::ii) :: 20 Word.word)) in + (let mappingpatterns_230 = ((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in (let op1 = (encdec_uop_backwards mappingpatterns_230) in - UTYPE (imm,rd,op1))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (p00 = (vec_of_bits [B1,B1,B0,B1,B1,B1,B1] :: 7 Word.word)))) then - (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in - (let (imm_18_13 :: 6 bits) = ((subrange_vec_dec v__175 (( 30 :: int)::ii) (( 25 :: int)::ii) :: 6 Word.word)) in - (let (imm_12_9 :: 4 bits) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 21 :: int)::ii) :: 4 Word.word)) in - (let (imm_8 :: 1 bits) = ((subrange_vec_dec v__175 (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word)) in - (let (imm_7_0 :: 8 bits) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 12 :: int)::ii) :: 8 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + UTYPE (imm,rd,op1)))))) + else if (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B1,B1,B0,B1,B1,B1,B1] :: 7 Word.word)))) then + (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm_8 :: 1 bits) = ((subrange_vec_dec v__177 (( 20 :: int)::ii) (( 20 :: int)::ii) :: 1 Word.word)) in + (let (imm_7_0 :: 8 bits) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 12 :: int)::ii) :: 8 Word.word)) in + (let (imm_19 :: 1 bits) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in + (let (imm_18_13 :: 6 bits) = ((subrange_vec_dec v__177 (( 30 :: int)::ii) (( 25 :: int)::ii) :: 6 Word.word)) in + (let (imm_12_9 :: 4 bits) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 21 :: int)::ii) :: 4 Word.word)) in RISCV_JAL ((concat_vec imm_19 ((concat_vec imm_7_0 ((concat_vec imm_8 @@ -22936,24 +23310,26 @@ definition encdec_backwards :: "(32)Word.word \<Rightarrow> ast " where :: 11 Word.word)) :: 12 Word.word)) :: 20 Word.word)) - :: 21 Word.word),rd))))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((p10 = (vec_of_bits [B1,B1,B0,B0,B1,B1,B1] :: 7 Word.word)))) \<and> (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))))))) then - (let (imm :: 12 Word.word) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - RISCV_JALR (imm,rs1,rd)))) - else if ((let mappingpatterns_240 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p00 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((encdec_bop_backwards_matches mappingpatterns_240)) \<and> (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B1,B1] :: 7 Word.word)))))))) then - (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in - (let (imm7_5_0 :: 6 bits) = ((subrange_vec_dec v__175 (( 30 :: int)::ii) (( 25 :: int)::ii) :: 6 Word.word)) in - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_240 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let (imm5_4_1 :: 4 bits) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 8 :: int)::ii) :: 4 Word.word)) in - (let (imm5_0 :: 1 bits) = ((subrange_vec_dec v__175 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in + :: 21 Word.word),rd)))))))) + else if ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B1,B1] :: 7 Word.word))))))) then + (let (imm :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + RISCV_JALR (imm,rs1,rd))))) + else if (((((let mappingpatterns_240 = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + (((encdec_bop_backwards_matches mappingpatterns_240)) \<and> (if ((encdec_bop_backwards_matches mappingpatterns_240)) then + (let op1 = (encdec_bop_backwards mappingpatterns_240) in + True) + else False)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (imm7_6 :: 1 bits) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 31 :: int)::ii) :: 1 Word.word)) in + (let (imm7_5_0 :: 6 bits) = ((subrange_vec_dec v__177 (( 30 :: int)::ii) (( 25 :: int)::ii) :: 6 Word.word)) in + (let (imm5_4_1 :: 4 bits) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 8 :: int)::ii) :: 4 Word.word)) in + (let (imm5_0 :: 1 bits) = ((subrange_vec_dec v__177 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_240 = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in (let op1 = (encdec_bop_backwards mappingpatterns_240) in BTYPE ((concat_vec imm7_6 ((concat_vec imm5_0 @@ -22961,441 +23337,393 @@ definition encdec_backwards :: "(32)Word.word \<Rightarrow> ast " where ((concat_vec imm5_4_1 (vec_of_bits [B0] :: 1 Word.word) :: 5 Word.word)) :: 11 Word.word)) :: 12 Word.word)) - :: 13 Word.word),rs2,rs1,op1))))))))) - else if ((let mappingpatterns_250 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p00 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((encdec_iop_backwards_matches mappingpatterns_250)) \<and> (((p00 = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))))))) then - (let (imm :: 12 Word.word) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_250 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + :: 13 Word.word),rs2,rs1,op1)))))))))) + else if (((((let mappingpatterns_250 = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + (((encdec_iop_backwards_matches mappingpatterns_250)) \<and> (if ((encdec_iop_backwards_matches mappingpatterns_250)) then + (let op1 = (encdec_iop_backwards mappingpatterns_250) in + True) + else False)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let (imm :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let mappingpatterns_250 = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in (let op1 = (encdec_iop_backwards mappingpatterns_250) in - ITYPE (imm,rs1,rd,op1)))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))))) then - (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + ITYPE (imm,rs1,rd,op1))))))) + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in SHIFTIOP (shamt,rs1,rd,RISCV_SLLI)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))))) then - (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in SHIFTIOP (shamt,rs1,rd,RISCV_SRLI)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0] :: 6 Word.word))))))))) then - (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0] :: 6 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 6 Word.word) = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 20 :: int)::ii) :: 6 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in SHIFTIOP (shamt,rs1,rd,RISCV_SRAI)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_ADD)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_SUB)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_SLL)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B1,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_SLT)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B1,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_SLTU)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_XOR)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_SRL)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_SRA)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B1,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_OR)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B1,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPE (rs2,rs1,rd,RISCV_AND)))) - else if ((let mappingpatterns_260 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_270 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p00 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((size_bits_backwards_matches mappingpatterns_270)) \<and> ((bool_bits_backwards_matches mappingpatterns_260))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B1] :: 7 Word.word))))))))) then - (let (imm :: 12 Word.word) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_260 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_270 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if (((((let mappingpatterns_270 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_260 = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_270)) \<and> (if ((size_bits_backwards_matches mappingpatterns_270)) then + (let size1 = (size_bits_backwards mappingpatterns_270) in + (((bool_bits_backwards_matches mappingpatterns_260)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_260)) then + (let is_unsigned = (bool_bits_backwards mappingpatterns_260) in + True) + else False))) + else False))))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let (imm :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let mappingpatterns_270 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_260 = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_270) in (let is_unsigned = (bool_bits_backwards mappingpatterns_260) in - LOAD (imm,rs1,rd,is_unsigned,size1,False,False)))))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_280 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((size_bits_backwards_matches mappingpatterns_280)) \<and> (((p10 = (vec_of_bits [B0,B1,B0,B0,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0] :: 1 Word.word))))))))) then - (let (imm7 :: 7 bits) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_280 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let (imm5 :: 5 bits) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + LOAD (imm,rs1,rd,is_unsigned,size1,False,False))))))))) + else if (((((let mappingpatterns_280 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_280)) \<and> (if ((size_bits_backwards_matches mappingpatterns_280)) then + (let size1 = (size_bits_backwards mappingpatterns_280) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (imm7 :: 7 bits) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (imm7 :: 7 bits) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in + (let (imm5 :: 5 bits) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_280 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_280) in - STORE ((concat_vec imm7 imm5 :: 12 Word.word),rs2,rs1,size1,False,False))))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((p10 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))))))) then - (let (imm :: 12 Word.word) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - ADDIW (imm,rs1,rd)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + STORE ((concat_vec imm7 imm5 :: 12 Word.word),rs2,rs1,size1,False,False)))))))) + else if ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word))))))) then + (let (imm :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + ADDIW (imm,rs1,rd))))) + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in SHIFTW (shamt,rs1,rd,RISCV_SLLI)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in SHIFTW (shamt,rs1,rd,RISCV_SRLI)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in SHIFTW (shamt,rs1,rd,RISCV_SRAI)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPEW (rs2,rs1,rd,RISCV_ADDW)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPEW (rs2,rs1,rd,RISCV_SUBW)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPEW (rs2,rs1,rd,RISCV_SLLW)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPEW (rs2,rs1,rd,RISCV_SRLW)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in RTYPEW (rs2,rs1,rd,RISCV_SRAW)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let (mappingpatterns_290 :: 3 bits) = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((encdec_mul_op_backwards_matches mappingpatterns_290)) \<and> (((p10 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (mappingpatterns_290 :: 3 bits) = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + SHIFTIWOP (shamt,rs1,rd,RISCV_SLLIW)))) + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + SHIFTIWOP (shamt,rs1,rd,RISCV_SRLIW)))) + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (shamt :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + SHIFTIWOP (shamt,rs1,rd,RISCV_SRAIW)))) + else if (((((let (mappingpatterns_290 :: 3 bits) = + ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + (((encdec_mul_op_backwards_matches mappingpatterns_290)) \<and> (if ((encdec_mul_op_backwards_matches mappingpatterns_290)) then + (let (high, signed1, signed2) = (encdec_mul_op_backwards mappingpatterns_290) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (mappingpatterns_290 :: 3 bits) = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in (let (high, signed1, signed2) = (encdec_mul_op_backwards mappingpatterns_290) in MUL (rs2,rs1,rd,high,signed1,signed2)))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) in - (let mappingpatterns_300 = ((subrange_vec_dec v__175 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((bool_not_bits_backwards_matches mappingpatterns_300)) \<and> (((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B1,B0] :: 2 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_300 = ((subrange_vec_dec v__175 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if (((((let mappingpatterns_300 = ((subrange_vec_dec v__177 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (((bool_not_bits_backwards_matches mappingpatterns_300)) \<and> (if ((bool_not_bits_backwards_matches mappingpatterns_300)) then + (let s = (bool_not_bits_backwards mappingpatterns_300) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_300 = ((subrange_vec_dec v__177 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let s = (bool_not_bits_backwards mappingpatterns_300) in DIV (rs2,rs1,rd,s)))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) in - (let mappingpatterns_310 = ((subrange_vec_dec v__175 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((bool_not_bits_backwards_matches mappingpatterns_310)) \<and> (((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B1,B1] :: 2 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_310 = ((subrange_vec_dec v__175 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if (((((let mappingpatterns_310 = ((subrange_vec_dec v__177 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (((bool_not_bits_backwards_matches mappingpatterns_310)) \<and> (if ((bool_not_bits_backwards_matches mappingpatterns_310)) then + (let s = (bool_not_bits_backwards mappingpatterns_310) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_310 = ((subrange_vec_dec v__177 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let s = (bool_not_bits_backwards mappingpatterns_310) in REM (rs2,rs1,rd,s)))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in MULW (rs2,rs1,rd)))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) in - (let mappingpatterns_320 = ((subrange_vec_dec v__175 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((bool_not_bits_backwards_matches mappingpatterns_320)) \<and> (((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B1,B0] :: 2 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_320 = ((subrange_vec_dec v__175 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if (((((let mappingpatterns_320 = ((subrange_vec_dec v__177 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (((bool_not_bits_backwards_matches mappingpatterns_320)) \<and> (if ((bool_not_bits_backwards_matches mappingpatterns_320)) then + (let s = (bool_not_bits_backwards mappingpatterns_320) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word))))))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_320 = ((subrange_vec_dec v__177 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let s = (bool_not_bits_backwards mappingpatterns_320) in DIVW (rs2,rs1,rd,s)))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) in - (let mappingpatterns_330 = ((subrange_vec_dec v__175 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((bool_not_bits_backwards_matches mappingpatterns_330)) \<and> (((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B1,B1] :: 2 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_330 = ((subrange_vec_dec v__175 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if (((((let mappingpatterns_330 = ((subrange_vec_dec v__177 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (((bool_not_bits_backwards_matches mappingpatterns_330)) \<and> (if ((bool_not_bits_backwards_matches mappingpatterns_330)) then + (let s = (bool_not_bits_backwards mappingpatterns_330) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word))))))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_330 = ((subrange_vec_dec v__177 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let s = (bool_not_bits_backwards mappingpatterns_330) in REMW (rs2,rs1,rd,s)))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B0,B0,B0,B1,B1,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0] :: 4 Word.word))))))))))) then - (let (pred :: 4 Word.word) = ((subrange_vec_dec v__175 (( 27 :: int)::ii) (( 24 :: int)::ii) :: 4 Word.word)) in - (let (succ :: 4 Word.word) = ((subrange_vec_dec v__175 (( 23 :: int)::ii) (( 20 :: int)::ii) :: 4 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) = (vec_of_bits [B0,B0,B0,B0] :: 4 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 0 :: int)::ii) :: 20 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B1] + :: 20 Word.word))))))) then + (let (succ :: 4 Word.word) = ((subrange_vec_dec v__177 (( 23 :: int)::ii) (( 20 :: int)::ii) :: 4 Word.word)) in + (let (pred :: 4 Word.word) = ((subrange_vec_dec v__177 (( 27 :: int)::ii) (( 24 :: int)::ii) :: 4 Word.word)) in FENCE (pred,succ))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B0,B0,B0,B1,B1,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word))))))))))) then + else if (((v__177 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0, + B0,B0,B0,B0,B1,B1,B1,B1] + :: 32 Word.word)))) then FENCEI () - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word))))))))))) then + else if (((v__177 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then ECALL () - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p50 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((((((((((((p50 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p40)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p30 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p20)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0] :: 7 Word.word)))))))))))) then + else if (((v__177 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then MRET () - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p50 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((((((((((((p50 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p40)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p30 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p20)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0] :: 7 Word.word)))))))))))) then + else if (((v__177 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then SRET () - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word))))))))))) then + else if (((v__177 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then EBREAK () - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word))))))))))) then + else if (((v__177 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then WFI () - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((((((p30 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p20)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B1] :: 7 Word.word)))))))))) then - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + else if ((((((((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B0,B0,B1] :: 7 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 0 :: int)::ii) :: 15 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B0,B0,B1,B1] :: 15 Word.word))))))) + then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in SFENCE_VMA (rs1,rs2))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_340 = ((subrange_vec_dec v__175 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_350 = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_360 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((((((((((((((size_bits_backwards_matches mappingpatterns_360)) \<and> ((bool_bits_backwards_matches mappingpatterns_350))))) \<and> ((bool_bits_backwards_matches mappingpatterns_340))))) \<and> (((p30 = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word))))))) \<and> (((p20 = (vec_of_bits [B0] :: 1 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word))))))))))))))) then - (let mappingpatterns_340 = ((subrange_vec_dec v__175 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_350 = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_360 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if (((((let mappingpatterns_360 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_350 = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_340 = ((subrange_vec_dec v__177 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_360)) \<and> (if ((size_bits_backwards_matches mappingpatterns_360)) then + (let size1 = (size_bits_backwards mappingpatterns_360) in + (((bool_bits_backwards_matches mappingpatterns_350)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_350)) then + (let rl = (bool_bits_backwards mappingpatterns_350) in + (((bool_bits_backwards_matches mappingpatterns_340)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_340)) then + (let aq = (bool_bits_backwards mappingpatterns_340) in + True) + else False))) + else False))) + else False)))))) \<and> ((((((((regbits_to_regno ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)))) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word)))))) \<and> ((((((((regbits_to_regno + ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)))) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word)))))))))))))))) then + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_360 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_350 = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_340 = ((subrange_vec_dec v__177 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_360) in (let rl = (bool_bits_backwards mappingpatterns_350) in (let aq = (bool_bits_backwards mappingpatterns_340) in LOADRES (aq,rl,rs1,size1,rd))))))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_370 = ((subrange_vec_dec v__175 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_380 = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_390 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((((size_bits_backwards_matches mappingpatterns_390)) \<and> ((bool_bits_backwards_matches mappingpatterns_380))))) \<and> ((bool_bits_backwards_matches mappingpatterns_370))))) \<and> (((p20 = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B0] :: 1 Word.word))))))) \<and> (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word)))))))))))))) then - (let mappingpatterns_370 = ((subrange_vec_dec v__175 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_380 = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_390 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if (((((let mappingpatterns_390 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_380 = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_370 = ((subrange_vec_dec v__177 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_390)) \<and> (if ((size_bits_backwards_matches mappingpatterns_390)) then + (let size1 = (size_bits_backwards mappingpatterns_390) in + (((bool_bits_backwards_matches mappingpatterns_380)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_380)) then + (let rl = (bool_bits_backwards mappingpatterns_380) in + (((bool_bits_backwards_matches mappingpatterns_370)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_370)) then + (let aq = (bool_bits_backwards mappingpatterns_370) in + True) + else False))) + else False))) + else False)))))) \<and> ((((((((regbits_to_regno ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)))) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word)))))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word))))))))))))) then + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_390 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_380 = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_370 = ((subrange_vec_dec v__177 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_390) in (let rl = (bool_bits_backwards mappingpatterns_380) in (let aq = (bool_bits_backwards mappingpatterns_370) in STORECON (aq,rl,rs2,rs1,size1,rd)))))))))) - else if ((let mappingpatterns_400 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_410 = ((subrange_vec_dec v__175 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_420 = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let p00 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_430 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((((size_bits_backwards_matches mappingpatterns_430)) \<and> ((bool_bits_backwards_matches mappingpatterns_420))))) \<and> ((bool_bits_backwards_matches mappingpatterns_410))))) \<and> ((encdec_amoop_backwards_matches mappingpatterns_400))))) \<and> (((p10 = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0] :: 1 Word.word)))))))))))) then - (let mappingpatterns_400 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_410 = ((subrange_vec_dec v__175 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_420 = ((subrange_vec_dec v__175 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_430 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + else if (((((let mappingpatterns_400 = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_430 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_420 = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_410 = ((subrange_vec_dec v__177 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_400 = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_430)) \<and> (if ((size_bits_backwards_matches mappingpatterns_430)) then + (let size1 = (size_bits_backwards mappingpatterns_430) in + (((bool_bits_backwards_matches mappingpatterns_420)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_420)) then + (let rl = (bool_bits_backwards mappingpatterns_420) in + (((bool_bits_backwards_matches mappingpatterns_410)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_410)) then + (let aq = (bool_bits_backwards mappingpatterns_410) in + (((encdec_amoop_backwards_matches mappingpatterns_400)) \<and> (if ((encdec_amoop_backwards_matches mappingpatterns_400)) then + (let op1 = (encdec_amoop_backwards mappingpatterns_400) in + True) + else False))) + else False))) + else False))) + else False)))))))) \<and> ((((((((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word)))))))))) then + (let mappingpatterns_400 = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in + (let (rs2 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_430 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_420 = ((subrange_vec_dec v__177 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_410 = ((subrange_vec_dec v__177 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_400 = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_430) in (let rl = (bool_bits_backwards mappingpatterns_420) in (let aq = (bool_bits_backwards mappingpatterns_410) in (let op1 = (encdec_amoop_backwards mappingpatterns_400) in - AMO (op1,aq,rl,rs2,rs1,size1,rd)))))))))))) - else if ((let mappingpatterns_440 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_450 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p00 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((encdec_csrop_backwards_matches mappingpatterns_450)) \<and> ((bool_bits_backwards_matches mappingpatterns_440))))) \<and> (((p00 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))))) then - (let (csr :: 12 Word.word) = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__175 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_440 = ((subrange_vec_dec v__175 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_450 = ((subrange_vec_dec v__175 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let (rd :: 5 Word.word) = ((subrange_vec_dec v__175 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + AMO (op1,aq,rl,rs2,rs1,size1,rd))))))))))))) + else if (((((let mappingpatterns_450 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_440 = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in + (((encdec_csrop_backwards_matches mappingpatterns_450)) \<and> (if ((encdec_csrop_backwards_matches mappingpatterns_450)) then + (let op1 = (encdec_csrop_backwards mappingpatterns_450) in + (((bool_bits_backwards_matches mappingpatterns_440)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_440)) then + (let is_imm = (bool_bits_backwards mappingpatterns_440) in + True) + else False))) + else False))))) \<and> (((((subrange_vec_dec v__177 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let (csr :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let (rs1 :: 5 Word.word) = ((subrange_vec_dec v__177 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: 5 Word.word) = ((subrange_vec_dec v__177 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (csr :: 12 Word.word) = ((subrange_vec_dec v__177 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let mappingpatterns_450 = ((subrange_vec_dec v__177 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_440 = ((subrange_vec_dec v__177 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in (let op1 = (encdec_csrop_backwards mappingpatterns_450) in (let is_imm = (bool_bits_backwards mappingpatterns_440) in - CSR (csr,rs1,rd,is_imm,op1)))))))) - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 15 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in - (let p50 = ((subrange_vec_dec v__175 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) in - ((((((((((((((((p50 = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((p40 = (vec_of_bits [B0,B1,B0] :: 3 Word.word))))))) \<and> (((p30 = (vec_of_bits [B0,B0] :: 2 Word.word))))))) \<and> (((p20 = (vec_of_bits [B0] :: 1 Word.word))))))) \<and> (((p10 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))))))) \<and> (((p00 = (vec_of_bits [B1,B1,B1,B1,B1,B0,B1,B0,B1,B1,B0,B1,B1,B1,B1,B0] :: 16 Word.word)))))))))))) - then + CSR (csr,rs1,rd,is_imm,op1))))))))) + else if (((v__177 = (vec_of_bits [B1,B1,B1,B1,B1,B0,B1,B0,B1,B1,B0,B1,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B1,B0,B1,B1] + :: 32 Word.word)))) then STOP_FETCHING () - else if ((let p00 = ((subrange_vec_dec v__175 (( 31 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) in - (let p10 = ((subrange_vec_dec v__175 (( 15 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) in - (let p20 = ((subrange_vec_dec v__175 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in - (let p30 = ((subrange_vec_dec v__175 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in - (let p40 = ((subrange_vec_dec v__175 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in - (let p50 = ((subrange_vec_dec v__175 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) in - ((((((((((((((((p50 = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((p40 = (vec_of_bits [B0,B1,B0] :: 3 Word.word))))))) \<and> (((p30 = (vec_of_bits [B0,B0] :: 2 Word.word))))))) \<and> (((p20 = (vec_of_bits [B0] :: 1 Word.word))))))) \<and> (((p10 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))))))) \<and> (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B1,B1,B0,B1,B1,B1,B1,B0] :: 16 Word.word)))))))))))) - then + else if (((v__177 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B1,B1,B0,B1,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B1,B0,B1,B1] + :: 32 Word.word)))) then THREAD_START () - else ILLEGAL v__175))" + else ILLEGAL v__177))" (*val encdec_forwards_matches : ast -> bool*) fun encdec_forwards_matches :: " ast \<Rightarrow> bool " where " encdec_forwards_matches (UTYPE (imm,rd,op1)) = ( True )" -|" encdec_forwards_matches (RISCV_JAL (v__224,rd)) = ( - if ((let p00 = ((subrange_vec_dec v__224 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in - (p00 = (vec_of_bits [B0] :: 1 Word.word)))) then +|" encdec_forwards_matches (RISCV_JAL (v__391,rd)) = ( + if (((((subrange_vec_dec v__391 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) then True else - (let g__17 = (RISCV_JAL (v__224,rd)) in + (let g__13 = (RISCV_JAL (v__391,rd)) in False))" |" encdec_forwards_matches (RISCV_JALR (imm,rs1,rd)) = ( True )" -|" encdec_forwards_matches (BTYPE (v__225,rs2,rs1,op1)) = ( - if ((let p00 = ((subrange_vec_dec v__225 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) in - (p00 = (vec_of_bits [B0] :: 1 Word.word)))) then +|" encdec_forwards_matches (BTYPE (v__393,rs2,rs1,op1)) = ( + if (((((subrange_vec_dec v__393 (( 0 :: int)::ii) (( 0 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) then True else - (let g__17 = (BTYPE (v__225,rs2,rs1,op1)) in + (let g__13 = (BTYPE (v__393,rs2,rs1,op1)) in False))" |" encdec_forwards_matches (ITYPE (imm,rs1,rd,op1)) = ( True )" |" encdec_forwards_matches (SHIFTIOP (shamt,rs1,rd,RISCV_SLLI)) = ( True )" @@ -23412,7 +23740,7 @@ fun encdec_forwards_matches :: " ast \<Rightarrow> bool " where |" encdec_forwards_matches (RTYPE (rs2,rs1,rd,RISCV_OR)) = ( True )" |" encdec_forwards_matches (RTYPE (rs2,rs1,rd,RISCV_AND)) = ( True )" |" encdec_forwards_matches (LOAD (imm,rs1,rd,is_unsigned,size1,False,False)) = ( True )" -|" encdec_forwards_matches (STORE (v__226,rs2,rs1,size1,False,False)) = ( True )" +|" encdec_forwards_matches (STORE (v__395,rs2,rs1,size1,False,False)) = ( True )" |" encdec_forwards_matches (ADDIW (imm,rs1,rd)) = ( True )" |" encdec_forwards_matches (SHIFTW (shamt,rs1,rd,RISCV_SLLI)) = ( True )" |" encdec_forwards_matches (SHIFTW (shamt,rs1,rd,RISCV_SRLI)) = ( True )" @@ -23422,6 +23750,9 @@ fun encdec_forwards_matches :: " ast \<Rightarrow> bool " where |" encdec_forwards_matches (RTYPEW (rs2,rs1,rd,RISCV_SLLW)) = ( True )" |" encdec_forwards_matches (RTYPEW (rs2,rs1,rd,RISCV_SRLW)) = ( True )" |" encdec_forwards_matches (RTYPEW (rs2,rs1,rd,RISCV_SRAW)) = ( True )" +|" encdec_forwards_matches (SHIFTIWOP (shamt,rs1,rd,RISCV_SLLIW)) = ( True )" +|" encdec_forwards_matches (SHIFTIWOP (shamt,rs1,rd,RISCV_SRLIW)) = ( True )" +|" encdec_forwards_matches (SHIFTIWOP (shamt,rs1,rd,RISCV_SRAIW)) = ( True )" |" encdec_forwards_matches (MUL (rs2,rs1,rd,high,signed1,signed2)) = ( True )" |" encdec_forwards_matches (DIV (rs2,rs1,rd,s)) = ( True )" |" encdec_forwards_matches (REM (rs2,rs1,rd,s)) = ( True )" @@ -23443,455 +23774,397 @@ fun encdec_forwards_matches :: " ast \<Rightarrow> bool " where |" encdec_forwards_matches (STOP_FETCHING (_)) = ( True )" |" encdec_forwards_matches (THREAD_START (_)) = ( True )" |" encdec_forwards_matches (ILLEGAL (s)) = ( True )" -|" encdec_forwards_matches g__17 = ( False )" +|" encdec_forwards_matches g__13 = ( False )" (*val encdec_backwards_matches : mword ty32 -> bool*) definition encdec_backwards_matches :: "(32)Word.word \<Rightarrow> bool " where " encdec_backwards_matches arg0 = ( - (let v__227 = arg0 in - if ((let mappingpatterns_00 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - encdec_uop_backwards_matches mappingpatterns_00)) then - (let mappingpatterns_00 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in + (let v__396 = arg0 in + if ((let mappingpatterns_00 = ((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in + (((encdec_uop_backwards_matches mappingpatterns_00)) \<and> (if ((encdec_uop_backwards_matches mappingpatterns_00)) then + (let op1 = (encdec_uop_backwards mappingpatterns_00) in + True) + else False)))) then + (let mappingpatterns_00 = ((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in (let op1 = (encdec_uop_backwards mappingpatterns_00) in True)) - else if ((let p00 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (p00 = (vec_of_bits [B1,B1,B0,B1,B1,B1,B1] :: 7 Word.word)))) then + else if (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B1,B1,B0,B1,B1,B1,B1] :: 7 Word.word)))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((p10 = (vec_of_bits [B1,B1,B0,B0,B1,B1,B1] :: 7 Word.word)))) \<and> (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))))))) then + else if ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B1,B1,B1] :: 7 Word.word))))))) then True - else if ((let mappingpatterns_10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p00 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((encdec_bop_backwards_matches mappingpatterns_10)) \<and> (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B1,B1] :: 7 Word.word)))))))) then - (let mappingpatterns_10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + else if (((((let mappingpatterns_10 = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + (((encdec_bop_backwards_matches mappingpatterns_10)) \<and> (if ((encdec_bop_backwards_matches mappingpatterns_10)) then + (let op1 = (encdec_bop_backwards mappingpatterns_10) in + True) + else False)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B1,B1,B0,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let mappingpatterns_10 = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in (let op1 = (encdec_bop_backwards mappingpatterns_10) in True)) - else if ((let mappingpatterns_20 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p00 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((encdec_iop_backwards_matches mappingpatterns_20)) \<and> (((p00 = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))))))) then - (let mappingpatterns_20 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + else if (((((let mappingpatterns_20 = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + (((encdec_iop_backwards_matches mappingpatterns_20)) \<and> (if ((encdec_iop_backwards_matches mappingpatterns_20)) then + (let op1 = (encdec_iop_backwards mappingpatterns_20) in + True) + else False)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let mappingpatterns_20 = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in (let op1 = (encdec_iop_backwards mappingpatterns_20) in True)) - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0] :: 6 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0] :: 6 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 26 :: int)::ii) :: 6 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0] :: 6 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B1,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B1,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B1,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B1,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let mappingpatterns_30 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_40 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p00 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((size_bits_backwards_matches mappingpatterns_40)) \<and> ((bool_bits_backwards_matches mappingpatterns_30))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B1,B1] :: 7 Word.word))))))))) then - (let mappingpatterns_30 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_40 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + else if (((((let mappingpatterns_40 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_30 = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_40)) \<and> (if ((size_bits_backwards_matches mappingpatterns_40)) then + (let size1 = (size_bits_backwards mappingpatterns_40) in + (((bool_bits_backwards_matches mappingpatterns_30)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_30)) then + (let is_unsigned = (bool_bits_backwards mappingpatterns_30) in + True) + else False))) + else False))))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let mappingpatterns_40 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_30 = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_40) in (let is_unsigned = (bool_bits_backwards mappingpatterns_30) in True)))) - else if ((let p00 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_50 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((size_bits_backwards_matches mappingpatterns_50)) \<and> (((p10 = (vec_of_bits [B0,B1,B0,B0,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0] :: 1 Word.word))))))))) then - (let mappingpatterns_50 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + else if (((((let mappingpatterns_50 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_50)) \<and> (if ((size_bits_backwards_matches mappingpatterns_50)) then + (let size1 = (size_bits_backwards mappingpatterns_50) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let mappingpatterns_50 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_50) in True)) - else if ((let p00 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((p10 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p00 = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))))))) then + else if ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word))))))) then + True + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + True + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B1,B0,B1] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let (mappingpatterns_60 :: 3 bits) = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((encdec_mul_op_backwards_matches mappingpatterns_60)) \<and> (((p10 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word))))))))) then - (let (mappingpatterns_60 :: 3 bits) = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B0,B0,B0,B0] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B1,B0,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then + True + else if (((((let (mappingpatterns_60 :: 3 bits) = + ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in + (((encdec_mul_op_backwards_matches mappingpatterns_60)) \<and> (if ((encdec_mul_op_backwards_matches mappingpatterns_60)) then + (let (high, signed1, signed2) = (encdec_mul_op_backwards mappingpatterns_60) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))))))))) then + (let (mappingpatterns_60 :: 3 bits) = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in (let (high, signed1, signed2) = (encdec_mul_op_backwards mappingpatterns_60) in True)) - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) in - (let mappingpatterns_70 = ((subrange_vec_dec v__227 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((bool_not_bits_backwards_matches mappingpatterns_70)) \<and> (((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B1,B0] :: 2 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))))))))) then - (let mappingpatterns_70 = ((subrange_vec_dec v__227 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + else if (((((let mappingpatterns_70 = ((subrange_vec_dec v__396 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (((bool_not_bits_backwards_matches mappingpatterns_70)) \<and> (if ((bool_not_bits_backwards_matches mappingpatterns_70)) then + (let s = (bool_not_bits_backwards mappingpatterns_70) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))))))))) then + (let mappingpatterns_70 = ((subrange_vec_dec v__396 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let s = (bool_not_bits_backwards mappingpatterns_70) in True)) - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) in - (let mappingpatterns_80 = ((subrange_vec_dec v__227 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((bool_not_bits_backwards_matches mappingpatterns_80)) \<and> (((p20 = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B1,B1] :: 2 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))))))))) then - (let mappingpatterns_80 = ((subrange_vec_dec v__227 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + else if (((((let mappingpatterns_80 = ((subrange_vec_dec v__396 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (((bool_not_bits_backwards_matches mappingpatterns_80)) \<and> (if ((bool_not_bits_backwards_matches mappingpatterns_80)) then + (let s = (bool_not_bits_backwards mappingpatterns_80) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))))))))) then + (let mappingpatterns_80 = ((subrange_vec_dec v__396 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let s = (bool_not_bits_backwards mappingpatterns_80) in True)) - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word)))))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) in - (let mappingpatterns_90 = ((subrange_vec_dec v__227 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((bool_not_bits_backwards_matches mappingpatterns_90)) \<and> (((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B1,B0] :: 2 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))))))))) then - (let mappingpatterns_90 = ((subrange_vec_dec v__227 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + else if (((((let mappingpatterns_90 = ((subrange_vec_dec v__396 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (((bool_not_bits_backwards_matches mappingpatterns_90)) \<and> (if ((bool_not_bits_backwards_matches mappingpatterns_90)) then + (let s = (bool_not_bits_backwards mappingpatterns_90) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word))))))))))))) then + (let mappingpatterns_90 = ((subrange_vec_dec v__396 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let s = (bool_not_bits_backwards mappingpatterns_90) in True)) - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) in - (let mappingpatterns_100 = ((subrange_vec_dec v__227 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((bool_not_bits_backwards_matches mappingpatterns_100)) \<and> (((p20 = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B1,B1] :: 2 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))))))))) then - (let mappingpatterns_100 = ((subrange_vec_dec v__227 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + else if (((((let mappingpatterns_100 = ((subrange_vec_dec v__396 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in + (((bool_not_bits_backwards_matches mappingpatterns_100)) \<and> (if ((bool_not_bits_backwards_matches mappingpatterns_100)) then + (let s = (bool_not_bits_backwards mappingpatterns_100) in + True) + else False)))) \<and> ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B1] :: 7 Word.word)))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 13 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B1,B1,B0,B1,B1] :: 7 Word.word))))))))))))) then + (let mappingpatterns_100 = ((subrange_vec_dec v__396 (( 12 :: int)::ii) (( 12 :: int)::ii) :: 1 Word.word)) in (let s = (bool_not_bits_backwards mappingpatterns_100) in True)) - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B0,B0,B0,B1,B1,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0] :: 4 Word.word))))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 28 :: int)::ii) :: 4 Word.word)) = (vec_of_bits [B0,B0,B0,B0] :: 4 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 19 :: int)::ii) (( 0 :: int)::ii) :: 20 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B1] + :: 20 Word.word))))))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B0,B0,B0,B1,B1,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B1] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word))))))))))) then + else if (((v__396 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0, + B0,B0,B0,B0,B1,B1,B1,B1] + :: 32 Word.word)))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] :: 12 Word.word))))))))))) then + else if (((v__396 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p50 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((((((((((((p50 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p40)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p30 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p20)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0] :: 7 Word.word)))))))))))) then + else if (((v__396 = (vec_of_bits [B0,B0,B1,B1,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p50 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((((((((((((p50 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p40)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p30 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p20)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0] :: 7 Word.word)))))))))))) then + else if (((v__396 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1] :: 12 Word.word))))))))))) then + else if (((v__396 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((p40 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p30)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p20 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1] :: 12 Word.word))))))))))) then + else if (((v__396 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B0,B0,B0,B1,B0,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B1,B1,B1,B0,B0,B1,B1] + :: 32 Word.word)))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((((((p30 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word)))) \<and> (((((regbits_to_regno p20)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((p10 = (vec_of_bits [B0,B0,B0] :: 3 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0,B0,B0,B1,B0,B0,B1] :: 7 Word.word)))))))))) then + else if ((((((((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 25 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B1,B0,B0,B1] :: 7 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 0 :: int)::ii) :: 15 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B1,B1,B1,B0,B0,B1,B1] :: 15 Word.word))))))) + then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_110 = ((subrange_vec_dec v__227 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_120 = ((subrange_vec_dec v__227 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_130 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((((((((((((((size_bits_backwards_matches mappingpatterns_130)) \<and> ((bool_bits_backwards_matches mappingpatterns_120))))) \<and> ((bool_bits_backwards_matches mappingpatterns_110))))) \<and> (((p30 = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word))))))) \<and> (((p20 = (vec_of_bits [B0] :: 1 Word.word))))))) \<and> (((((regbits_to_regno p10)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word))))))))) \<and> (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word))))))))))))))) then - (let mappingpatterns_110 = ((subrange_vec_dec v__227 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_120 = ((subrange_vec_dec v__227 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_130 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + else if (((((let mappingpatterns_130 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_120 = ((subrange_vec_dec v__396 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_110 = ((subrange_vec_dec v__396 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_130)) \<and> (if ((size_bits_backwards_matches mappingpatterns_130)) then + (let size1 = (size_bits_backwards mappingpatterns_130) in + (((bool_bits_backwards_matches mappingpatterns_120)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_120)) then + (let rl = (bool_bits_backwards mappingpatterns_120) in + (((bool_bits_backwards_matches mappingpatterns_110)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_110)) then + (let aq = (bool_bits_backwards mappingpatterns_110) in + True) + else False))) + else False))) + else False)))))) \<and> ((((((((regbits_to_regno ((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)))) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B0] :: 5 Word.word)))))) \<and> ((((((((regbits_to_regno + ((subrange_vec_dec v__396 (( 24 :: int)::ii) (( 20 :: int)::ii) :: 5 Word.word)))) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B0,B0] :: 5 Word.word)))))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word)))))))))))))))) then + (let mappingpatterns_130 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_120 = ((subrange_vec_dec v__396 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_110 = ((subrange_vec_dec v__396 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_130) in (let rl = (bool_bits_backwards mappingpatterns_120) in (let aq = (bool_bits_backwards mappingpatterns_110) in True)))))) - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_140 = ((subrange_vec_dec v__227 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_150 = ((subrange_vec_dec v__227 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_160 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((((size_bits_backwards_matches mappingpatterns_160)) \<and> ((bool_bits_backwards_matches mappingpatterns_150))))) \<and> ((bool_bits_backwards_matches mappingpatterns_140))))) \<and> (((p20 = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word))))))) \<and> (((p10 = (vec_of_bits [B0] :: 1 Word.word))))))) \<and> (((((regbits_to_regno p00)) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word)))))))))))))) then - (let mappingpatterns_140 = ((subrange_vec_dec v__227 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_150 = ((subrange_vec_dec v__227 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_160 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + else if (((((let mappingpatterns_160 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_150 = ((subrange_vec_dec v__396 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_140 = ((subrange_vec_dec v__396 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_160)) \<and> (if ((size_bits_backwards_matches mappingpatterns_160)) then + (let size1 = (size_bits_backwards mappingpatterns_160) in + (((bool_bits_backwards_matches mappingpatterns_150)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_150)) then + (let rl = (bool_bits_backwards mappingpatterns_150) in + (((bool_bits_backwards_matches mappingpatterns_140)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_140)) then + (let aq = (bool_bits_backwards mappingpatterns_140) in + True) + else False))) + else False))) + else False)))))) \<and> ((((((((regbits_to_regno ((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)))) = ((regbits_to_regno (vec_of_bits [B0,B0,B0,B1,B1] :: 5 Word.word)))))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word))))))))))))) then + (let mappingpatterns_160 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_150 = ((subrange_vec_dec v__396 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_140 = ((subrange_vec_dec v__396 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_160) in (let rl = (bool_bits_backwards mappingpatterns_150) in (let aq = (bool_bits_backwards mappingpatterns_140) in True)))))) - else if ((let mappingpatterns_170 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_180 = ((subrange_vec_dec v__227 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_190 = ((subrange_vec_dec v__227 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let p00 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_200 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - (((((((((((((((size_bits_backwards_matches mappingpatterns_200)) \<and> ((bool_bits_backwards_matches mappingpatterns_190))))) \<and> ((bool_bits_backwards_matches mappingpatterns_180))))) \<and> ((encdec_amoop_backwards_matches mappingpatterns_170))))) \<and> (((p10 = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word))))))) \<and> (((p00 = (vec_of_bits [B0] :: 1 Word.word)))))))))))) then - (let mappingpatterns_170 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in - (let mappingpatterns_180 = ((subrange_vec_dec v__227 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_190 = ((subrange_vec_dec v__227 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_200 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + else if (((((let mappingpatterns_170 = ((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_200 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_190 = ((subrange_vec_dec v__396 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_180 = ((subrange_vec_dec v__396 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_170 = ((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in + (((size_bits_backwards_matches mappingpatterns_200)) \<and> (if ((size_bits_backwards_matches mappingpatterns_200)) then + (let size1 = (size_bits_backwards mappingpatterns_200) in + (((bool_bits_backwards_matches mappingpatterns_190)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_190)) then + (let rl = (bool_bits_backwards mappingpatterns_190) in + (((bool_bits_backwards_matches mappingpatterns_180)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_180)) then + (let aq = (bool_bits_backwards mappingpatterns_180) in + (((encdec_amoop_backwards_matches mappingpatterns_170)) \<and> (if ((encdec_amoop_backwards_matches mappingpatterns_170)) then + (let op1 = (encdec_amoop_backwards mappingpatterns_170) in + True) + else False))) + else False))) + else False))) + else False)))))))) \<and> ((((((((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) = (vec_of_bits [B0] :: 1 Word.word)))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B1,B0,B1,B1,B1,B1] :: 7 Word.word)))))))))) then + (let mappingpatterns_170 = ((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in + (let mappingpatterns_200 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_190 = ((subrange_vec_dec v__396 (( 25 :: int)::ii) (( 25 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_180 = ((subrange_vec_dec v__396 (( 26 :: int)::ii) (( 26 :: int)::ii) :: 1 Word.word)) in + (let mappingpatterns_170 = ((subrange_vec_dec v__396 (( 31 :: int)::ii) (( 27 :: int)::ii) :: 5 Word.word)) in (let size1 = (size_bits_backwards mappingpatterns_200) in (let rl = (bool_bits_backwards mappingpatterns_190) in (let aq = (bool_bits_backwards mappingpatterns_180) in (let op1 = (encdec_amoop_backwards mappingpatterns_170) in - True)))))))) - else if ((let mappingpatterns_210 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_220 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in - (let p00 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) in - ((((((encdec_csrop_backwards_matches mappingpatterns_220)) \<and> ((bool_bits_backwards_matches mappingpatterns_210))))) \<and> (((p00 = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))))) then - (let mappingpatterns_210 = ((subrange_vec_dec v__227 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in - (let mappingpatterns_220 = ((subrange_vec_dec v__227 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + True))))))))) + else if (((((let mappingpatterns_220 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_210 = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in + (((encdec_csrop_backwards_matches mappingpatterns_220)) \<and> (if ((encdec_csrop_backwards_matches mappingpatterns_220)) then + (let op1 = (encdec_csrop_backwards mappingpatterns_220) in + (((bool_bits_backwards_matches mappingpatterns_210)) \<and> (if ((bool_bits_backwards_matches mappingpatterns_210)) then + (let is_imm = (bool_bits_backwards mappingpatterns_210) in + True) + else False))) + else False))))) \<and> (((((subrange_vec_dec v__396 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B1,B1,B1,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let mappingpatterns_220 = ((subrange_vec_dec v__396 (( 13 :: int)::ii) (( 12 :: int)::ii) :: 2 Word.word)) in + (let mappingpatterns_210 = ((subrange_vec_dec v__396 (( 14 :: int)::ii) (( 14 :: int)::ii) :: 1 Word.word)) in (let op1 = (encdec_csrop_backwards mappingpatterns_220) in (let is_imm = (bool_bits_backwards mappingpatterns_210) in True)))) - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 15 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in - (let p50 = ((subrange_vec_dec v__227 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) in - ((((((((((((((((p50 = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((p40 = (vec_of_bits [B0,B1,B0] :: 3 Word.word))))))) \<and> (((p30 = (vec_of_bits [B0,B0] :: 2 Word.word))))))) \<and> (((p20 = (vec_of_bits [B0] :: 1 Word.word))))))) \<and> (((p10 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))))))) \<and> (((p00 = (vec_of_bits [B1,B1,B1,B1,B1,B0,B1,B0,B1,B1,B0,B1,B1,B1,B1,B0] :: 16 Word.word)))))))))))) - then + else if (((v__396 = (vec_of_bits [B1,B1,B1,B1,B1,B0,B1,B0,B1,B1,B0,B1,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B1,B0,B1,B1] + :: 32 Word.word)))) then True - else if ((let p00 = ((subrange_vec_dec v__227 (( 31 :: int)::ii) (( 16 :: int)::ii) :: 16 Word.word)) in - (let p10 = ((subrange_vec_dec v__227 (( 15 :: int)::ii) (( 8 :: int)::ii) :: 8 Word.word)) in - (let p20 = ((subrange_vec_dec v__227 (( 7 :: int)::ii) (( 7 :: int)::ii) :: 1 Word.word)) in - (let p30 = ((subrange_vec_dec v__227 (( 6 :: int)::ii) (( 5 :: int)::ii) :: 2 Word.word)) in - (let p40 = ((subrange_vec_dec v__227 (( 4 :: int)::ii) (( 2 :: int)::ii) :: 3 Word.word)) in - (let p50 = ((subrange_vec_dec v__227 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) in - ((((((((((((((((p50 = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((p40 = (vec_of_bits [B0,B1,B0] :: 3 Word.word))))))) \<and> (((p30 = (vec_of_bits [B0,B0] :: 2 Word.word))))))) \<and> (((p20 = (vec_of_bits [B0] :: 1 Word.word))))))) \<and> (((p10 = (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0] :: 8 Word.word))))))) \<and> (((p00 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B1,B1,B0,B1,B1,B1,B1,B0] :: 16 Word.word)))))))))))) - then + else if (((v__396 = (vec_of_bits [B1,B1,B0,B0,B0,B0,B0,B0,B1,B1,B0,B1,B1,B1,B1,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B1,B0,B1,B1] + :: 32 Word.word)))) then True else True))" fun print_insn :: " ast \<Rightarrow> string " where - " print_insn (NOP (g__13)) = ( (''nop''))" + " print_insn (NOP (_)) = ( (''nop''))" |" print_insn (C_ADDI4SPN (rdc,nzimm)) = ( - (op@) (''c.addi4spn '') - (((op@) ((reg_name_abi ((creg2reg_bits rdc :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits nzimm)))))))" + (@) (''c.addi4spn '') + (((@) ((reg_name_abi ((creg2reg_bits rdc :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits nzimm)))))))" |" print_insn (C_LW (uimm,rsc,rdc)) = ( - (op@) (''c.lw '') - (((op@) ((reg_name_abi ((creg2reg_bits rdc :: 5 Word.word)))) - (((op@) ('', '') - (((op@) ((reg_name_abi ((creg2reg_bits rsc :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits uimm)))))))))))" + (@) (''c.lw '') + (((@) ((reg_name_abi ((creg2reg_bits rdc :: 5 Word.word)))) + (((@) ('', '') + (((@) ((reg_name_abi ((creg2reg_bits rsc :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits uimm)))))))))))" |" print_insn (C_LD (uimm,rsc,rdc)) = ( - (op@) (''c.ld '') - (((op@) ((reg_name_abi ((creg2reg_bits rdc :: 5 Word.word)))) - (((op@) ('', '') - (((op@) ((reg_name_abi ((creg2reg_bits rsc :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits uimm)))))))))))" + (@) (''c.ld '') + (((@) ((reg_name_abi ((creg2reg_bits rdc :: 5 Word.word)))) + (((@) ('', '') + (((@) ((reg_name_abi ((creg2reg_bits rsc :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits uimm)))))))))))" |" print_insn (C_SW (uimm,rsc1,rsc2)) = ( - (op@) (''c.sw '') - (((op@) ((reg_name_abi ((creg2reg_bits rsc1 :: 5 Word.word)))) - (((op@) ('', '') - (((op@) ((reg_name_abi ((creg2reg_bits rsc2 :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits uimm)))))))))))" + (@) (''c.sw '') + (((@) ((reg_name_abi ((creg2reg_bits rsc1 :: 5 Word.word)))) + (((@) ('', '') + (((@) ((reg_name_abi ((creg2reg_bits rsc2 :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits uimm)))))))))))" |" print_insn (C_SD (uimm,rsc1,rsc2)) = ( - (op@) (''c.sd '') - (((op@) ((reg_name_abi ((creg2reg_bits rsc1 :: 5 Word.word)))) - (((op@) ('', '') - (((op@) ((reg_name_abi ((creg2reg_bits rsc2 :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits uimm)))))))))))" + (@) (''c.sd '') + (((@) ((reg_name_abi ((creg2reg_bits rsc1 :: 5 Word.word)))) + (((@) ('', '') + (((@) ((reg_name_abi ((creg2reg_bits rsc2 :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits uimm)))))))))))" |" print_insn (C_ADDI (nzi,rsd)) = ( - (op@) (''c.addi '') - (((op@) ((reg_name_abi rsd)) (((op@) ('', '') ((string_of_bits nzi)))))))" -|" print_insn (C_JAL (imm)) = ( (op@) (''c.jal '') ((string_of_bits imm)))" + (@) (''c.addi '') + (((@) ((reg_name_abi rsd)) (((@) ('', '') ((string_of_bits nzi)))))))" +|" print_insn (C_JAL (imm)) = ( (@) (''c.jal '') ((string_of_bits imm)))" |" print_insn (C_ADDIW (imm,rsd)) = ( - (op@) (''c.addiw '') - (((op@) ((reg_name_abi rsd)) (((op@) ('', '') ((string_of_bits imm)))))))" + (@) (''c.addiw '') + (((@) ((reg_name_abi rsd)) (((@) ('', '') ((string_of_bits imm)))))))" |" print_insn (C_LI (imm,rd)) = ( - (op@) (''c.li '') - (((op@) ((reg_name_abi rd)) (((op@) ('', '') ((string_of_bits imm)))))))" -|" print_insn (C_ADDI16SP (imm)) = ( (op@) (''c.addi16sp '') ((string_of_bits imm)))" + (@) (''c.li '') + (((@) ((reg_name_abi rd)) (((@) ('', '') ((string_of_bits imm)))))))" +|" print_insn (C_ADDI16SP (imm)) = ( (@) (''c.addi16sp '') ((string_of_bits imm)))" |" print_insn (C_LUI (imm,rd)) = ( - (op@) (''c.lui '') - (((op@) ((reg_name_abi rd)) (((op@) ('', '') ((string_of_bits imm)))))))" + (@) (''c.lui '') + (((@) ((reg_name_abi rd)) (((@) ('', '') ((string_of_bits imm)))))))" |" print_insn (C_SRLI (shamt,rsd)) = ( - (op@) (''c.srli '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits shamt)))))))" + (@) (''c.srli '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits shamt)))))))" |" print_insn (C_SRAI (shamt,rsd)) = ( - (op@) (''c.srai '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits shamt)))))))" + (@) (''c.srai '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits shamt)))))))" |" print_insn (C_ANDI (imm,rsd)) = ( - (op@) (''c.andi '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits imm)))))))" + (@) (''c.andi '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits imm)))))))" |" print_insn (C_SUB (rsd,rs2)) = ( - (op@) (''c.sub '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" + (@) (''c.sub '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" |" print_insn (C_XOR (rsd,rs2)) = ( - (op@) (''c.xor '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" + (@) (''c.xor '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" |" print_insn (C_OR (rsd,rs2)) = ( - (op@) (''c.or '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" + (@) (''c.or '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" |" print_insn (C_AND (rsd,rs2)) = ( - (op@) (''c.and '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" + (@) (''c.and '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" |" print_insn (C_SUBW (rsd,rs2)) = ( - (op@) (''c.subw '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" + (@) (''c.subw '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" |" print_insn (C_ADDW (rsd,rs2)) = ( - (op@) (''c.addw '') - (((op@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) - (((op@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" -|" print_insn (C_J (imm)) = ( (op@) (''c.j '') ((string_of_bits imm)))" + (@) (''c.addw '') + (((@) ((reg_name_abi ((creg2reg_bits rsd :: 5 Word.word)))) + (((@) ('', '') ((reg_name_abi ((creg2reg_bits rs2 :: 5 Word.word)))))))))" +|" print_insn (C_J (imm)) = ( (@) (''c.j '') ((string_of_bits imm)))" |" print_insn (C_BEQZ (imm,rs)) = ( - (op@) (''c.beqz '') - (((op@) ((reg_name_abi ((creg2reg_bits rs :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits imm)))))))" + (@) (''c.beqz '') + (((@) ((reg_name_abi ((creg2reg_bits rs :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits imm)))))))" |" print_insn (C_BNEZ (imm,rs)) = ( - (op@) (''c.bnez '') - (((op@) ((reg_name_abi ((creg2reg_bits rs :: 5 Word.word)))) - (((op@) ('', '') ((string_of_bits imm)))))))" + (@) (''c.bnez '') + (((@) ((reg_name_abi ((creg2reg_bits rs :: 5 Word.word)))) + (((@) ('', '') ((string_of_bits imm)))))))" |" print_insn (C_SLLI (shamt,rsd)) = ( - (op@) (''c.slli '') - (((op@) ((reg_name_abi rsd)) (((op@) ('', '') ((string_of_bits shamt)))))))" + (@) (''c.slli '') + (((@) ((reg_name_abi rsd)) (((@) ('', '') ((string_of_bits shamt)))))))" |" print_insn (C_LWSP (uimm,rd)) = ( - (op@) (''c.lwsp '') - (((op@) ((reg_name_abi rd)) (((op@) ('', '') ((string_of_bits uimm)))))))" + (@) (''c.lwsp '') + (((@) ((reg_name_abi rd)) (((@) ('', '') ((string_of_bits uimm)))))))" |" print_insn (C_LDSP (uimm,rd)) = ( - (op@) (''c.ldsp '') - (((op@) ((reg_name_abi rd)) (((op@) ('', '') ((string_of_bits uimm)))))))" + (@) (''c.ldsp '') + (((@) ((reg_name_abi rd)) (((@) ('', '') ((string_of_bits uimm)))))))" |" print_insn (C_SWSP (uimm,rd)) = ( - (op@) (''c.swsp '') - (((op@) ((reg_name_abi rd)) (((op@) ('', '') ((string_of_bits uimm)))))))" + (@) (''c.swsp '') + (((@) ((reg_name_abi rd)) (((@) ('', '') ((string_of_bits uimm)))))))" |" print_insn (C_SDSP (uimm,rd)) = ( - (op@) (''c.sdsp '') - (((op@) ((reg_name_abi rd)) (((op@) ('', '') ((string_of_bits uimm)))))))" -|" print_insn (C_JR (rs1)) = ( (op@) (''c.jr '') ((reg_name_abi rs1)))" -|" print_insn (C_JALR (rs1)) = ( (op@) (''c.jalr '') ((reg_name_abi rs1)))" + (@) (''c.sdsp '') + (((@) ((reg_name_abi rd)) (((@) ('', '') ((string_of_bits uimm)))))))" +|" print_insn (C_JR (rs1)) = ( (@) (''c.jr '') ((reg_name_abi rs1)))" +|" print_insn (C_JALR (rs1)) = ( (@) (''c.jalr '') ((reg_name_abi rs1)))" |" print_insn (C_MV (rd,rs2)) = ( - (op@) (''c.mv '') - (((op@) ((reg_name_abi rd)) (((op@) ('', '') ((reg_name_abi rs2)))))))" + (@) (''c.mv '') + (((@) ((reg_name_abi rd)) (((@) ('', '') ((reg_name_abi rs2)))))))" |" print_insn (C_ADD (rsd,rs2)) = ( - (op@) (''c.add '') - (((op@) ((reg_name_abi rsd)) (((op@) ('', '') ((reg_name_abi rs2)))))))" -|" print_insn (STOP_FETCHING (g__14)) = ( (''stop_fetching''))" -|" print_insn (THREAD_START (g__15)) = ( (''thread_start''))" -|" print_insn (ILLEGAL (s)) = ( (op@) (''illegal '') ((string_of_bits s)))" -|" print_insn (C_ILLEGAL (g__16)) = ( (''c.illegal''))" + (@) (''c.add '') + (((@) ((reg_name_abi rsd)) (((@) ('', '') ((reg_name_abi rs2)))))))" +|" print_insn (STOP_FETCHING (_)) = ( (''stop_fetching''))" +|" print_insn (THREAD_START (_)) = ( (''thread_start''))" +|" print_insn (ILLEGAL (s)) = ( (@) (''illegal '') ((string_of_bits s)))" +|" print_insn (C_ILLEGAL (_)) = ( (''c.illegal''))" |" print_insn insn = ( assembly_forwards insn )" @@ -23918,25 +24191,25 @@ definition fetch :: " unit \<Rightarrow>((register_value),(FetchResult),(except return (((((cast_unit_vec0 ((access_vec_dec w__1 (( 1 :: int)::ii))) :: 1 Word.word)) \<noteq> (vec_of_bits [B0] :: 1 Word.word)))))) (haveRVC () \<bind> (\<lambda> (w__2 :: bool) . return ((\<not> w__2))))) \<bind> (\<lambda> (w__4 :: bool) . if w__4 then - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__5 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__5 :: xlenbits) . return (F_Error (E_Fetch_Addr_Align,w__5))) else - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__6 :: xlenbits) . translateAddr w__6 Execute Instruction \<bind> (\<lambda> (w__7 :: TR_Result) . (case w__7 of TR_Failure (e) => - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__8 :: 64 Word.word) . return (F_Error (e,w__8))) + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__8 :: xlenbits) . return (F_Error (e,w__8))) | TR_Address (ppclo) => (checked_mem_read Instruction ppclo (( 2 :: int)::ii) :: ( ( 16 Word.word)MemoryOpResult) M) \<bind> (\<lambda> (w__9 :: ( 16 Word.word) MemoryOpResult) . (case w__9 of MemException (e) => - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__10 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__10 :: xlenbits) . return (F_Error (E_Fetch_Access_Fault,w__10))) | MemValue (ilo) => if ((isRVC ilo)) then return (F_RVC ilo) else - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__11 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__11 :: xlenbits) . (let (PChi :: xlenbits) = ((add_vec_int w__11 (( 2 :: int)::ii) :: 64 Word.word)) in translateAddr PChi Execute Instruction \<bind> (\<lambda> (w__12 :: TR_Result) . (case w__12 of @@ -23977,30 +24250,30 @@ definition step :: " int \<Rightarrow>((register_value),(bool*bool),(exception) (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__7 :: xlenbits) . (let (_ :: unit) = (print_endline - (((op@) (''['') - (((op@) ((stringFromInteger step_no)) - (((op@) (''] ['') - (((op@) ((privLevel_to_str w__6)) - (((op@) ('']: '') - (((op@) ((string_of_bits w__7)) - (((op@) ('' ('') - (((op@) ((string_of_bits h)) ('') <no-decode>'')))))))))))))))))) in + (((@) (''['') + (((@) ((stringFromInteger step_no)) + (((@) (''] ['') + (((@) ((privLevel_to_str w__6)) + (((@) ('']: '') + (((@) ((string_of_bits w__7)) + (((@) ('' ('') + (((@) ((string_of_bits h)) ('') <no-decode>'')))))))))))))))))) in handle_decode_exception ((EXTZ (( 64 :: int)::ii) h :: 64 Word.word)) \<then> return (False, True)))) | Some (ast) => read_reg cur_privilege_ref \<bind> (\<lambda> (w__8 :: Privilege) . (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__9 :: xlenbits) . (let (_ :: unit) = (print_endline - (((op@) (''['') - (((op@) ((stringFromInteger step_no)) - (((op@) (''] ['') - (((op@) ((privLevel_to_str w__8)) - (((op@) ('']: '') - (((op@) ((string_of_bits w__9)) - (((op@) ('' ('') - (((op@) ((string_of_bits h)) - (((op@) ('') '') ((print_insn ast))))))))))))))))))))) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__10 :: 64 Word.word) . + (((@) (''['') + (((@) ((stringFromInteger step_no)) + (((@) (''] ['') + (((@) ((privLevel_to_str w__8)) + (((@) ('']: '') + (((@) ((string_of_bits w__9)) + (((@) ('' ('') + (((@) ((string_of_bits h)) + (((@) ('') '') ((print_insn ast))))))))))))))))))))) in + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__10 :: xlenbits) . (write_reg nextPC_ref ((add_vec_int w__10 (( 2 :: int)::ii) :: 64 Word.word)) \<then> execute ast) \<bind> (\<lambda> (w__11 :: bool) . return (w__11, True)))))) ) @@ -24011,30 +24284,30 @@ definition step :: " int \<Rightarrow>((register_value),(bool*bool),(exception) (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__14 :: xlenbits) . (let (_ :: unit) = (print_endline - (((op@) (''['') - (((op@) ((stringFromInteger step_no)) - (((op@) (''] ['') - (((op@) ((privLevel_to_str w__13)) - (((op@) ('']: '') - (((op@) ((string_of_bits w__14)) - (((op@) ('' ('') - (((op@) ((string_of_bits w)) ('') <no-decode>'')))))))))))))))))) in + (((@) (''['') + (((@) ((stringFromInteger step_no)) + (((@) (''] ['') + (((@) ((privLevel_to_str w__13)) + (((@) ('']: '') + (((@) ((string_of_bits w__14)) + (((@) ('' ('') + (((@) ((string_of_bits w)) ('') <no-decode>'')))))))))))))))))) in handle_decode_exception ((EXTZ (( 64 :: int)::ii) w :: 64 Word.word)) \<then> return (False, True)))) | Some (ast) => read_reg cur_privilege_ref \<bind> (\<lambda> (w__15 :: Privilege) . (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__16 :: xlenbits) . (let (_ :: unit) = (print_endline - (((op@) (''['') - (((op@) ((stringFromInteger step_no)) - (((op@) (''] ['') - (((op@) ((privLevel_to_str w__15)) - (((op@) ('']: '') - (((op@) ((string_of_bits w__16)) - (((op@) ('' ('') - (((op@) ((string_of_bits w)) - (((op@) ('') '') ((print_insn ast))))))))))))))))))))) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: 64 Word.word) . + (((@) (''['') + (((@) ((stringFromInteger step_no)) + (((@) (''] ['') + (((@) ((privLevel_to_str w__15)) + (((@) ('']: '') + (((@) ((string_of_bits w__16)) + (((@) ('' ('') + (((@) ((string_of_bits w)) + (((@) ('') '') ((print_insn ast))))))))))))))))))))) in + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__17 :: xlenbits) . (write_reg nextPC_ref ((add_vec_int w__17 (( 4 :: int)::ii) :: 64 Word.word)) \<then> execute ast) \<bind> (\<lambda> (w__18 :: bool) . return (w__18, True)))))) ) @@ -24226,7 +24499,7 @@ fun num_of_trans_kind :: " trans_kind \<Rightarrow> int " where definition GPRstr :: "(string)list " where " GPRstr = ( [(''x31''),(''x30''),(''x29''),(''x28''),(''x27''),(''x26''),(''x25''),(''x24''),(''x23''),(''x22''),(''x21''),(''x20''),(''x19''),(''x18''),(''x17''),(''x16''), - (''x15''),(''x14''),(''x13''),(''x12''),(''x21''),(''x10''),(''x9''),(''x8''),(''x7''),(''x6''),(''x5''),(''x4''),(''x3''),(''x2''),(''x1''),(''x0'')])" + (''x15''),(''x14''),(''x13''),(''x12''),(''x11''),(''x10''),(''x9''),(''x8''),(''x7''),(''x6''),(''x5''),(''x4''),(''x3''),(''x2''),(''x1''),(''x0'')])" definition CIA_fp :: " regfp " where @@ -24259,7 +24532,7 @@ definition initial_analysis :: " ast \<Rightarrow>((register_value),((regfp)lis (if (((((regbits_to_regno rd)) = (( 0 :: int)::ii)))) then oR else (RFull ((access_list_dec GPRstr ((regbits_to_regno rd))))) # oR) in (let (offset :: 64 bits) = ((EXTS (( 64 :: int)::ii) imm :: 64 Word.word)) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: xlenbits) . (let (Nias :: niafps) = ([NIAFP_concrete_address ((add_vec w__0 offset :: 64 Word.word))]) in (let (ik :: instruction_kind) = (IK_branch () ) in return (Nias, aR, iR, ik, oR)))))) @@ -24283,7 +24556,7 @@ definition initial_analysis :: " ast \<Rightarrow>((register_value),((regfp)lis else (RFull ((access_list_dec GPRstr ((regbits_to_regno rs1))))) # iR) in (let ik = (IK_branch () ) in (let (offset :: 64 bits) = ((EXTS (( 64 :: int)::ii) imm :: 64 Word.word)) in - (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: 64 Word.word) . + (read_reg PC_ref :: ( 64 Word.word) M) \<bind> (\<lambda> (w__1 :: xlenbits) . (let (Nias :: niafps) = ([NIAFP_concrete_address ((add_vec w__1 offset :: 64 Word.word)),NIAFP_successor () ]) in return (Nias, aR, iR, ik, oR))))))) @@ -24390,33 +24663,33 @@ definition initial_analysis :: " ast \<Rightarrow>((register_value),((regfp)lis return (Nias, aR, iR, ik, oR)))) | FENCE (pred,succ) => (case (pred, succ) of - (v__276, v__277) => - if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word))))))) then + (v__610, v__611) => + if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_rw_rw) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_r_rw) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_r_r) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_rw_w) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_w_w) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_w_rw) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_rw_r) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_r_w) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B1] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B1,B0] :: 2 Word.word))))))) then return (IK_barrier Barrier_RISCV_w_r) - else if ((((((((subrange_vec_dec v__276 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__277 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word))))))) then + else if ((((((((subrange_vec_dec v__610 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word)))) \<and> (((((subrange_vec_dec v__611 (( 1 :: int)::ii) (( 0 :: int)::ii) :: 2 Word.word)) = (vec_of_bits [B0,B0] :: 2 Word.word))))))) then return (IK_simple () ) else internal_error (''barrier type not implemented in initial_analysis'') ) \<bind> (\<lambda> (w__17 :: instruction_kind) . (let (ik :: instruction_kind) = w__17 in return (Nias, aR, iR, ik, oR))) | FENCEI (_) => - (let (ik :: instruction_kind) = (IK_barrier Barrier_RISCV_i) in + (let (ik :: instruction_kind) = (IK_simple () ) in return (Nias, aR, iR, ik, oR)) | LOADRES (aq,rl,rs1,width,rd) => (let (iR :: regfps) = diff --git a/snapshots/isabelle/riscv/Riscv_duopod.thy b/snapshots/isabelle/riscv/Riscv_duopod.thy new file mode 100644 index 00000000..70926709 --- /dev/null +++ b/snapshots/isabelle/riscv/Riscv_duopod.thy @@ -0,0 +1,539 @@ +chapter \<open>Generated by Lem from \<open>riscv_duopod.lem\<close>.\<close> + +theory "Riscv_duopod" + +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_duopod_types" + "Riscv_extras" + +begin + +(*Generated by Sail from riscv_duopod.*) +(*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_duopod_types*) +(*open import Riscv_extras*) + +(*val spc_forwards : unit -> string*) + +definition spc_forwards :: " unit \<Rightarrow> string " where + " spc_forwards _ = ( ('' ''))" + + +(*val spc_backwards : string -> unit*) + +definition spc_backwards :: " string \<Rightarrow> unit " where + " spc_backwards s = ( () )" + + +(*val opt_spc_forwards : unit -> string*) + +definition opt_spc_forwards :: " unit \<Rightarrow> string " where + " opt_spc_forwards _ = ( (''''))" + + +(*val opt_spc_backwards : string -> unit*) + +definition opt_spc_backwards :: " string \<Rightarrow> unit " where + " opt_spc_backwards s = ( () )" + + +(*val def_spc_forwards : unit -> string*) + +definition def_spc_forwards :: " unit \<Rightarrow> string " where + " def_spc_forwards _ = ( ('' ''))" + + +(*val def_spc_backwards : string -> unit*) + +definition def_spc_backwards :: " string \<Rightarrow> unit " where + " def_spc_backwards s = ( () )" + + + + + + + + + + +(*val builtin_and_vec : forall 'n. bits 'n -> bits 'n -> bits 'n*) + + + +(*val builtin_or_vec : forall 'n. bits 'n -> bits 'n -> bits 'n*) + + + +(*val __raw_SetSlice_int : forall 'w. integer -> ii -> ii -> bits 'w -> ii*) + +(*val __GetSlice_int : forall 'n . Size 'n => integer -> ii -> ii -> mword 'n*) + +definition GetSlice_int :: " int \<Rightarrow> int \<Rightarrow> int \<Rightarrow>('n::len)Word.word " where + " GetSlice_int n m o1 = ( (get_slice_int n m o1 :: ( 'n::len)Word.word))" + + +(*val __raw_SetSlice_bits : forall 'n 'w. integer -> integer -> bits 'n -> ii -> bits 'w -> bits 'n*) + +(*val __raw_GetSlice_bits : forall 'n 'w . integer -> integer -> bits 'n -> ii -> bits 'w*) + +(*val cast_unit_vec : bitU -> mword ty1*) + +fun cast_unit_vec0 :: " bitU \<Rightarrow>(1)Word.word " where + " cast_unit_vec0 B0 = ( (vec_of_bits [B0] :: 1 Word.word))" +|" cast_unit_vec0 B1 = ( (vec_of_bits [B1] :: 1 Word.word))" + + +(*val DecStr : ii -> string*) + +(*val HexStr : ii -> string*) + +(*val __RISCV_write : forall 'int8_times_n. Size 'int8_times_n => mword ty64 -> integer -> mword 'int8_times_n -> M bool*) + +definition RISCV_write :: "(64)Word.word \<Rightarrow> int \<Rightarrow>('int8_times_n::len)Word.word \<Rightarrow>((register_value),(bool),(unit))monad " where + " RISCV_write addr width data = ( + write_ram (( 64 :: int)::ii) width + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word) addr data )" + + +(*val __TraceMemoryWrite : forall 'int8_times_n 'm. integer -> bits 'm -> bits 'int8_times_n -> unit*) + +(*val __RISCV_read : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> bool -> bool -> bool -> M (maybe (mword 'int8_times_n))*) + +fun RISCV_read :: "(64)Word.word \<Rightarrow> int \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow>((register_value),((('int8_times_n::len)Word.word)option),(unit))monad " where + " RISCV_read addr width False False False = ( + (MEMr (( 64 :: int)::ii) width + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word) addr + :: (( 'int8_times_n::len)Word.word) M) \<bind> (\<lambda> (w__0 :: ( 'int8_times_n::len)Word.word) . + return (Some w__0)))" +|" RISCV_read addr width True False False = ( + (MEMr_acquire (( 64 :: int)::ii) width + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word) addr + :: (( 'int8_times_n::len)Word.word) M) \<bind> (\<lambda> (w__1 :: ( 'int8_times_n::len)Word.word) . + return (Some w__1)))" +|" RISCV_read addr width True True False = ( + (MEMr_strong_acquire (( 64 :: int)::ii) width + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word) addr + :: (( 'int8_times_n::len)Word.word) M) \<bind> (\<lambda> (w__2 :: ( 'int8_times_n::len)Word.word) . + return (Some w__2)))" +|" RISCV_read addr width False False True = ( + (MEMr_reserved (( 64 :: int)::ii) width + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word) addr + :: (( 'int8_times_n::len)Word.word) M) \<bind> (\<lambda> (w__3 :: ( 'int8_times_n::len)Word.word) . + return (Some w__3)))" +|" RISCV_read addr width True False True = ( + (MEMr_reserved_acquire (( 64 :: int)::ii) width + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word) addr + :: (( 'int8_times_n::len)Word.word) M) \<bind> (\<lambda> (w__4 :: ( 'int8_times_n::len)Word.word) . + return (Some w__4)))" +|" RISCV_read addr width True True True = ( + (MEMr_reserved_strong_acquire (( 64 :: int)::ii) width + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word) addr + :: (( 'int8_times_n::len)Word.word) M) \<bind> (\<lambda> (w__5 :: ( 'int8_times_n::len)Word.word) . + return (Some w__5)))" +|" RISCV_read addr width False True False = ( return None )" +|" RISCV_read addr width False True True = ( return None )" + + +(*val __TraceMemoryRead : forall 'int8_times_n 'm. integer -> bits 'm -> bits 'int8_times_n -> unit*) + +(*val ex_nat : ii -> integer*) + +definition ex_nat :: " int \<Rightarrow> int " where + " ex_nat n = ( n )" + + +(*val ex_int : ii -> integer*) + +definition ex_int :: " int \<Rightarrow> int " where + " ex_int n = ( n )" + + +(*val coerce_int_nat : ii -> M ii*) + +definition coerce_int_nat :: " int \<Rightarrow>((register_value),(int),(unit))monad " where + " coerce_int_nat x = ( assert_exp True ('''') \<then> return x )" + + +(*val EXTS : forall 'n 'm . Size 'm, Size 'n => integer -> mword 'n -> mword 'm*) + +(*val EXTZ : forall 'n 'm . Size 'm, Size 'n => integer -> mword 'n -> mword 'm*) + +definition EXTS :: " int \<Rightarrow>('n::len)Word.word \<Rightarrow>('m::len)Word.word " where + " EXTS (m__tv :: int) v = ( (sign_extend v m__tv :: ( 'm::len)Word.word))" + + +definition EXTZ :: " int \<Rightarrow>('n::len)Word.word \<Rightarrow>('m::len)Word.word " where + " EXTZ (m__tv :: int) v = ( (zero_extend v m__tv :: ( 'm::len)Word.word))" + + +(*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 \<Rightarrow>('n::len)Word.word \<Rightarrow> bool " where + " zopz0zI_s x y = ( ((Word.sint x)) < ((Word.sint y)))" + + +definition zopz0zKzJ_s :: "('n::len)Word.word \<Rightarrow>('n::len)Word.word \<Rightarrow> bool " where + " zopz0zKzJ_s x y = ( ((Word.sint x)) \<ge> ((Word.sint y)))" + + +definition zopz0zI_u :: "('n::len)Word.word \<Rightarrow>('n::len)Word.word \<Rightarrow> bool " where + " zopz0zI_u x y = ( ((Word.uint x)) < ((Word.uint y)))" + + +definition zopz0zKzJ_u :: "('n::len)Word.word \<Rightarrow>('n::len)Word.word \<Rightarrow> bool " where + " zopz0zKzJ_u x y = ( ((Word.uint x)) \<ge> ((Word.uint y)))" + + +definition zopz0zIzJ_u :: "('n::len)Word.word \<Rightarrow>('n::len)Word.word \<Rightarrow> bool " where + " zopz0zIzJ_u x y = ( ((Word.uint x)) \<le> ((Word.uint y)))" + + +(*val bool_to_bits : bool -> mword ty1*) + +definition bool_to_bits :: " bool \<Rightarrow>(1)Word.word " where + " bool_to_bits x = ( if x then (vec_of_bits [B1] :: 1 Word.word) else (vec_of_bits [B0] :: 1 Word.word))" + + +(*val bit_to_bool : bitU -> bool*) + +fun bit_to_bool :: " bitU \<Rightarrow> bool " where + " bit_to_bool B1 = ( True )" +|" bit_to_bool B0 = ( False )" + + +(*val vector64 : ii -> mword ty64*) + +definition vector64 :: " int \<Rightarrow>(64)Word.word " where + " vector64 n = ( (get_slice_int (( 64 :: int)::ii) n (( 0 :: int)::ii) :: 64 Word.word))" + + +(*val to_bits : forall 'l . Size 'l => integer -> ii -> mword 'l*) + +definition to_bits :: " int \<Rightarrow> int \<Rightarrow>('l::len)Word.word " where + " to_bits l n = ( (get_slice_int l n (( 0 :: int)::ii) :: ( 'l::len)Word.word))" + + +(*val shift_right_arith64 : mword ty64 -> mword ty6 -> mword ty64*) + +definition shift_right_arith64 :: "(64)Word.word \<Rightarrow>(6)Word.word \<Rightarrow>(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)))" + + +(*val shift_right_arith32 : mword ty32 -> mword ty5 -> mword ty32*) + +definition shift_right_arith32 :: "(32)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(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)))" + + +(*val zeros : forall 'n . Size 'n => integer -> mword 'n*) + +definition zeros0 :: " int \<Rightarrow>('n::len)Word.word " where + " zeros0 n = ( (replicate_bits (vec_of_bits [B0] :: 1 Word.word) n :: ( 'n::len)Word.word))" + + +(*val regbits_to_regno : mword ty5 -> integer*) + +definition regbits_to_regno :: "(5)Word.word \<Rightarrow> int " where + " regbits_to_regno b = ( + (let r = (Word.uint b) in + r))" + + +(*val rX : integer -> M (mword ty64)*) + +definition rX :: " int \<Rightarrow>((register_value),((64)Word.word),(unit))monad " where + " rX l__0 = ( + if (((l__0 = (( 0 :: int)::ii)))) then + return (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word) + else + read_reg Xs_ref \<bind> (\<lambda> (w__0 :: xlen_t list) . + return ((access_list_dec w__0 l__0 :: 64 Word.word))))" + + +(*val wX : integer -> mword ty64 -> M unit*) + +definition wX :: " int \<Rightarrow>(64)Word.word \<Rightarrow>((register_value),(unit),(unit))monad " where + " wX r v = ( + if (((r \<noteq> (( 0 :: int)::ii)))) then + read_reg Xs_ref \<bind> (\<lambda> (w__0 :: xlen_t list) . + write_reg Xs_ref ((update_list_dec w__0 r v :: ( 64 Word.word) list))) + else return () )" + + +(*val MEMr : forall 'int8_times_n . Size 'int8_times_n => mword ty64 -> integer -> M (mword 'int8_times_n)*) + +definition MEMr0 :: "(64)Word.word \<Rightarrow> int \<Rightarrow>((register_value),(('int8_times_n::len)Word.word),(unit))monad " where + " MEMr0 addr width = ( + (RISCV_read addr width False False False ) \<bind> (\<lambda> w__0 . + return ((case w__0 of + Some (v) => v + | None => (zeros0 (((( 8 :: int)::ii) * width)) :: ( 'int8_times_n::len)Word.word) + ))))" + + +(*val iop_of_num : integer -> iop*) + +definition iop_of_num :: " int \<Rightarrow> iop " where + " iop_of_num arg0 = ( + (let p00 = arg0 in + if (((p00 = (( 0 :: int)::ii)))) then RISCV_ADDI + else if (((p00 = (( 1 :: int)::ii)))) then RISCV_SLTI + else if (((p00 = (( 2 :: int)::ii)))) then RISCV_SLTIU + else if (((p00 = (( 3 :: int)::ii)))) then RISCV_XORI + else if (((p00 = (( 4 :: int)::ii)))) then RISCV_ORI + else RISCV_ANDI))" + + +(*val num_of_iop : iop -> integer*) + +fun num_of_iop :: " iop \<Rightarrow> 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 decode : mword ty32 -> maybe ast*) + +(*val execute : ast -> M unit*) + +definition decode :: "(32)Word.word \<Rightarrow>(ast)option " where + " decode v__0 = ( + if ((((((((subrange_vec_dec v__0 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B0,B0] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B1,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let (imm :: 12 bits) = ((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let (rs1 :: regbits) = ((subrange_vec_dec v__0 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regbits) = ((subrange_vec_dec v__0 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm :: 12 bits) = ((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + Some (ITYPE (imm,rs1,rd,RISCV_ADDI)))))) + else if ((((((((subrange_vec_dec v__0 (( 14 :: int)::ii) (( 12 :: int)::ii) :: 3 Word.word)) = (vec_of_bits [B0,B1,B1] :: 3 Word.word)))) \<and> (((((subrange_vec_dec v__0 (( 6 :: int)::ii) (( 0 :: int)::ii) :: 7 Word.word)) = (vec_of_bits [B0,B0,B0,B0,B0,B1,B1] :: 7 Word.word))))))) then + (let (imm :: 12 bits) = ((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + (let (rs1 :: regbits) = ((subrange_vec_dec v__0 (( 19 :: int)::ii) (( 15 :: int)::ii) :: 5 Word.word)) in + (let (rd :: regbits) = ((subrange_vec_dec v__0 (( 11 :: int)::ii) (( 7 :: int)::ii) :: 5 Word.word)) in + (let (imm :: 12 bits) = ((subrange_vec_dec v__0 (( 31 :: int)::ii) (( 20 :: int)::ii) :: 12 Word.word)) in + Some (LOAD (imm,rs1,rd)))))) + else None )" + + +(*val execute_LOAD : mword ty12 -> mword ty5 -> mword ty5 -> M unit*) + +definition execute_LOAD :: "(12)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>((register_value),(unit),(unit))monad " where + " execute_LOAD imm rs1 rd = ( + (rX ((regbits_to_regno rs1)) :: ( 64 Word.word) M) \<bind> (\<lambda> (w__0 :: 64 Word.word) . + (let (addr :: xlen_t) = ((add_vec w__0 ((EXTS (( 64 :: int)::ii) imm :: 64 Word.word)) :: 64 Word.word)) in + (MEMr0 addr (( 8 :: int)::ii) :: ( 64 Word.word) M) \<bind> (\<lambda> (result :: xlen_t) . + wX ((regbits_to_regno rd)) result))))" + + +(*val execute_ITYPE : mword ty12 -> mword ty5 -> mword ty5 -> iop -> M unit*) + +fun execute_ITYPE :: "(12)Word.word \<Rightarrow>(5)Word.word \<Rightarrow>(5)Word.word \<Rightarrow> iop \<Rightarrow>((register_value),(unit),(unit))monad " where + " execute_ITYPE imm rs1 rd RISCV_ADDI = ( + (rX ((regbits_to_regno rs1)) :: ( 64 Word.word) M) \<bind> (\<lambda> rs1_val . + (let (imm_ext :: xlen_t) = ((EXTS (( 64 :: int)::ii) imm :: 64 Word.word)) in + (let result = ((add_vec rs1_val imm_ext :: 64 Word.word)) in + wX ((regbits_to_regno rd)) result))))" + + +fun execute :: " ast \<Rightarrow>((register_value),(unit),(unit))monad " where + " execute (ITYPE (imm,rs1,rd,arg3)) = ( execute_ITYPE imm rs1 rd arg3 )" +|" execute (LOAD (imm,rs1,rd)) = ( execute_LOAD imm rs1 rd )" + + +definition initial_regstate :: " regstate " where + " initial_regstate = ( + (| Xs = + ([(vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word), + (vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word)]), + nextPC = + ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word)), + PC = + ((vec_of_bits [B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0, + B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0,B0] + :: 64 Word.word)) |) )" + + + +end diff --git a/snapshots/isabelle/riscv/Riscv_duopod_lemmas.thy b/snapshots/isabelle/riscv/Riscv_duopod_lemmas.thy new file mode 100644 index 00000000..a4901d1c --- /dev/null +++ b/snapshots/isabelle/riscv/Riscv_duopod_lemmas.thy @@ -0,0 +1,61 @@ +theory Riscv_duopod_lemmas + imports + Sail.Sail2_values_lemmas + Sail.Sail2_state_lemmas + Riscv_duopod +begin + +abbreviation liftS ("\<lbrakk>_\<rbrakk>\<^sub>S") where "liftS \<equiv> liftState (get_regval, set_regval)" + +lemmas register_defs = get_regval_def set_regval_def Xs_ref_def nextPC_ref_def PC_ref_def + +lemma regval_vector_64_dec_bit[simp]: + "vector_64_dec_bit_of_regval (regval_of_vector_64_dec_bit v) = Some v" + by (auto simp: regval_of_vector_64_dec_bit_def) + +lemma vector_of_rv_rv_of_vector[simp]: + assumes "\<And>v. of_rv (rv_of v) = Some v" + shows "vector_of_regval of_rv (regval_of_vector rv_of len is_inc v) = Some v" +proof - + from assms have "of_rv \<circ> rv_of = Some" by auto + then show ?thesis by (auto simp: vector_of_regval_def regval_of_vector_def) +qed + +lemma option_of_rv_rv_of_option[simp]: + assumes "\<And>v. of_rv (rv_of v) = Some v" + shows "option_of_regval of_rv (regval_of_option rv_of v) = Some v" + using assms by (cases v) (auto simp: option_of_regval_def regval_of_option_def) + +lemma list_of_rv_rv_of_list[simp]: + assumes "\<And>v. of_rv (rv_of v) = Some v" + shows "list_of_regval of_rv (regval_of_list rv_of v) = Some v" +proof - + from assms have "of_rv \<circ> rv_of = Some" by auto + with assms show ?thesis by (induction v) (auto simp: list_of_regval_def regval_of_list_def) +qed + +lemma liftS_read_reg_Xs[liftState_simp]: + "\<lbrakk>read_reg Xs_ref\<rbrakk>\<^sub>S = readS (Xs \<circ> regstate)" + by (auto simp: liftState_read_reg_readS register_defs) + +lemma liftS_write_reg_Xs[liftState_simp]: + "\<lbrakk>write_reg Xs_ref v\<rbrakk>\<^sub>S = updateS (regstate_update (Xs_update (\<lambda>_. v)))" + by (auto simp: liftState_write_reg_updateS register_defs) + +lemma liftS_read_reg_nextPC[liftState_simp]: + "\<lbrakk>read_reg nextPC_ref\<rbrakk>\<^sub>S = readS (nextPC \<circ> regstate)" + by (auto simp: liftState_read_reg_readS register_defs) + +lemma liftS_write_reg_nextPC[liftState_simp]: + "\<lbrakk>write_reg nextPC_ref v\<rbrakk>\<^sub>S = updateS (regstate_update (nextPC_update (\<lambda>_. v)))" + by (auto simp: liftState_write_reg_updateS register_defs) + +lemma liftS_read_reg_PC[liftState_simp]: + "\<lbrakk>read_reg PC_ref\<rbrakk>\<^sub>S = readS (PC \<circ> regstate)" + by (auto simp: liftState_read_reg_readS register_defs) + +lemma liftS_write_reg_PC[liftState_simp]: + "\<lbrakk>write_reg PC_ref v\<rbrakk>\<^sub>S = updateS (regstate_update (PC_update (\<lambda>_. v)))" + by (auto simp: liftState_write_reg_updateS register_defs) + +end diff --git a/snapshots/isabelle/riscv/Riscv_duopod_types.thy b/snapshots/isabelle/riscv/Riscv_duopod_types.thy new file mode 100644 index 00000000..75a0b487 --- /dev/null +++ b/snapshots/isabelle/riscv/Riscv_duopod_types.thy @@ -0,0 +1,173 @@ +chapter \<open>Generated by Lem from \<open>riscv_duopod_types.lem\<close>.\<close> + +theory "Riscv_duopod_types" + +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" + +begin + +(*Generated by Sail from riscv_duopod.*) +(*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*) +type_synonym 'n bits =" ( 'n::len)Word.word " + + + +type_synonym xlen =" int " + +type_synonym xlen_t =" 64 bits " + +type_synonym 'n regno =" int " + +type_synonym regbits =" 5 bits " + +datatype iop = RISCV_ADDI | RISCV_SLTI | RISCV_SLTIU | RISCV_XORI | RISCV_ORI | RISCV_ANDI + + + +datatype ast = + ITYPE " (( 12 bits * regbits * regbits * iop))" | LOAD " (( 12 bits * regbits * regbits))" + + + +datatype register_value = + Regval_vector " ((ii * bool * register_value list))" + | Regval_list " ( register_value list)" + | Regval_option " ( register_value option)" + | Regval_vector_64_dec_bit " ( 64 Word.word)" + + + +record regstate = + Xs ::" ( 64 Word.word) list " + nextPC ::" 64 Word.word " + PC ::" 64 Word.word " + + + + + +(*val vector_64_dec_bit_of_regval : register_value -> maybe (mword ty64)*) + +fun vector_64_dec_bit_of_regval :: " register_value \<Rightarrow>((64)Word.word)option " where + " vector_64_dec_bit_of_regval (Regval_vector_64_dec_bit (v)) = ( Some v )" +|" vector_64_dec_bit_of_regval g__1 = ( None )" + + +(*val regval_of_vector_64_dec_bit : mword ty64 -> register_value*) + +definition regval_of_vector_64_dec_bit :: "(64)Word.word \<Rightarrow> register_value " where + " regval_of_vector_64_dec_bit v = ( Regval_vector_64_dec_bit v )" + + + + +(*val vector_of_regval : forall 'a. (register_value -> maybe 'a) -> register_value -> maybe (list 'a)*) +definition vector_of_regval :: "(register_value \<Rightarrow> 'a option)\<Rightarrow> register_value \<Rightarrow>('a list)option " where + " vector_of_regval of_regval1 = ( \<lambda>x . + (case x of + Regval_vector (_, _, v) => just_list (List.map of_regval1 v) + | _ => None + ) )" + + +(*val regval_of_vector : forall 'a. ('a -> register_value) -> integer -> bool -> list 'a -> register_value*) +definition regval_of_vector :: "('a \<Rightarrow> register_value)\<Rightarrow> int \<Rightarrow> bool \<Rightarrow> 'a list \<Rightarrow> register_value " where + " regval_of_vector regval_of1 size1 is_inc xs = ( Regval_vector (size1, is_inc, List.map regval_of1 xs))" + + +(*val list_of_regval : forall 'a. (register_value -> maybe 'a) -> register_value -> maybe (list 'a)*) +definition list_of_regval :: "(register_value \<Rightarrow> 'a option)\<Rightarrow> register_value \<Rightarrow>('a list)option " where + " list_of_regval of_regval1 = ( \<lambda>x . + (case x of + Regval_list v => just_list (List.map of_regval1 v) + | _ => None + ) )" + + +(*val regval_of_list : forall 'a. ('a -> register_value) -> list 'a -> register_value*) +definition regval_of_list :: "('a \<Rightarrow> register_value)\<Rightarrow> 'a list \<Rightarrow> register_value " where + " regval_of_list regval_of1 xs = ( Regval_list (List.map regval_of1 xs))" + + +(*val option_of_regval : forall 'a. (register_value -> maybe 'a) -> register_value -> maybe (maybe 'a)*) +definition option_of_regval :: "(register_value \<Rightarrow> 'a option)\<Rightarrow> register_value \<Rightarrow>('a option)option " where + " option_of_regval of_regval1 = ( \<lambda>x . + (case x of + Regval_option v => Some (Option.bind v of_regval1) + | _ => None + ) )" + + +(*val regval_of_option : forall 'a. ('a -> register_value) -> maybe 'a -> register_value*) +definition regval_of_option :: "('a \<Rightarrow> register_value)\<Rightarrow> 'a option \<Rightarrow> register_value " where + " regval_of_option regval_of1 v = ( Regval_option (map_option regval_of1 v))" + + + +definition Xs_ref :: "((regstate),(register_value),(((64)Word.word)list))register_ref " where + " Xs_ref = ( (| + name = (''Xs''), + read_from = (\<lambda> s . (Xs s)), + write_to = (\<lambda> v s . (( s (| Xs := v |)))), + of_regval = (\<lambda> v . vector_of_regval (\<lambda> v . vector_64_dec_bit_of_regval v) v), + regval_of = (\<lambda> v . regval_of_vector (\<lambda> v . regval_of_vector_64_dec_bit v)(( 32 :: int)) False v) |) )" + + +definition nextPC_ref :: "((regstate),(register_value),((64)Word.word))register_ref " where + " nextPC_ref = ( (| + name = (''nextPC''), + read_from = (\<lambda> s . (nextPC s)), + write_to = (\<lambda> v s . (( s (| nextPC := v |)))), + of_regval = (\<lambda> v . vector_64_dec_bit_of_regval v), + regval_of = (\<lambda> v . regval_of_vector_64_dec_bit v) |) )" + + +definition PC_ref :: "((regstate),(register_value),((64)Word.word))register_ref " where + " PC_ref = ( (| + name = (''PC''), + read_from = (\<lambda> s . (PC s)), + write_to = (\<lambda> v s . (( s (| PC := v |)))), + of_regval = (\<lambda> v . vector_64_dec_bit_of_regval v), + regval_of = (\<lambda> v . regval_of_vector_64_dec_bit v) |) )" + + +(*val get_regval : string -> regstate -> maybe register_value*) +definition get_regval :: " string \<Rightarrow> regstate \<Rightarrow>(register_value)option " where + " get_regval reg_name s = ( + if reg_name = (''Xs'') then Some ((regval_of Xs_ref) ((read_from Xs_ref) s)) else + if reg_name = (''nextPC'') then Some ((regval_of nextPC_ref) ((read_from nextPC_ref) s)) else + if reg_name = (''PC'') then Some ((regval_of PC_ref) ((read_from PC_ref) s)) else + None )" + + +(*val set_regval : string -> register_value -> regstate -> maybe regstate*) +definition set_regval :: " string \<Rightarrow> register_value \<Rightarrow> regstate \<Rightarrow>(regstate)option " where + " set_regval reg_name v s = ( + if reg_name = (''Xs'') then map_option (\<lambda> v . (write_to Xs_ref) v s) ((of_regval Xs_ref) v) else + if reg_name = (''nextPC'') then map_option (\<lambda> v . (write_to nextPC_ref) v s) ((of_regval nextPC_ref) v) else + if reg_name = (''PC'') then map_option (\<lambda> v . (write_to PC_ref) v s) ((of_regval PC_ref) v) else + None )" + + +definition register_accessors :: "(string \<Rightarrow> regstate \<Rightarrow>(register_value)option)*(string \<Rightarrow> register_value \<Rightarrow> regstate \<Rightarrow>(regstate)option)" where + " register_accessors = ( (get_regval, set_regval))" + + + +type_synonym( 'a, 'r) MR =" (register_value, regstate, 'a, 'r, unit) base_monadR " +type_synonym 'a M =" (register_value, regstate, 'a, unit) base_monad " +end diff --git a/snapshots/isabelle/riscv/Riscv_extras.thy b/snapshots/isabelle/riscv/Riscv_extras.thy index 8e371170..0a4528fe 100644 --- a/snapshots/isabelle/riscv/Riscv_extras.thy +++ b/snapshots/isabelle/riscv/Riscv_extras.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from riscv_extras.lem.\<close> +chapter \<open>Generated by Lem from \<open>riscv_extras.lem\<close>.\<close> theory "Riscv_extras" -imports - Main - "Lem_pervasives" - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_operators_mwords" - "Sail2_prompt_monad" - "Sail2_prompt" +imports + Main + "LEM.Lem_pervasives" + "LEM.Lem_pervasives_extra" + "Sail.Sail2_instr_kinds" + "Sail.Sail2_values" + "Sail.Sail2_operators_mwords" + "Sail.Sail2_prompt_monad" + "Sail.Sail2_prompt" begin @@ -122,10 +122,10 @@ definition MEMr_reserved_strong_acquire :: " int \<Rightarrow> int \<Rightarrow (*val write_ram : forall 'rv 'a 'b 'e. Size 'a, Size 'b => - integer -> integer -> bitvector 'a -> bitvector 'a -> bitvector 'b -> monad 'rv unit 'e*) -definition write_ram :: " int \<Rightarrow> int \<Rightarrow>('a::len)Word.word \<Rightarrow>('a::len)Word.word \<Rightarrow>('b::len)Word.word \<Rightarrow>('rv,(unit),'e)monad " where + integer -> integer -> bitvector 'a -> bitvector 'a -> bitvector 'b -> monad 'rv bool 'e*) +definition write_ram :: " int \<Rightarrow> int \<Rightarrow>('a::len)Word.word \<Rightarrow>('a::len)Word.word \<Rightarrow>('b::len)Word.word \<Rightarrow>('rv,(bool),'e)monad " where " write_ram addrsize size1 hexRAM address value1 = ( - write_mem_val instance_Sail2_values_Bitvector_Machine_word_mword_dict value1 \<bind> (\<lambda>x . (case x of _ => return () )) )" + write_mem_val instance_Sail2_values_Bitvector_Machine_word_mword_dict value1 )" (*val read_ram : forall 'rv 'a 'b 'e. Size 'a, Size 'b => diff --git a/snapshots/isabelle/riscv/Riscv_types.thy b/snapshots/isabelle/riscv/Riscv_types.thy index 78328442..ae9cd2e6 100644 --- a/snapshots/isabelle/riscv/Riscv_types.thy +++ b/snapshots/isabelle/riscv/Riscv_types.thy @@ -1,16 +1,16 @@ -chapter \<open>Generated by Lem from riscv_types.lem.\<close> +chapter \<open>Generated by Lem from \<open>riscv_types.lem\<close>.\<close> theory "Riscv_types" -imports - Main - "Lem_pervasives_extra" - "Sail2_instr_kinds" - "Sail2_values" - "Sail2_operators_mwords" - "Sail2_prompt_monad" - "Sail2_prompt" - "Sail2_string" +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" begin @@ -32,14 +32,14 @@ type_synonym half =" 16 bits " type_synonym word0 =" 32 bits " -type_synonym 'n regno =" int " - type_synonym regbits =" 5 bits " type_synonym cregbits =" 3 bits " type_synonym csreg =" 12 bits " +type_synonym 'n regno =" int " + type_synonym opcode =" 7 bits " type_synonym imm12 =" 12 bits " @@ -68,8 +68,25 @@ datatype ReadType = Instruction | Data +datatype word_width = BYTE | HALF | WORD | DOUBLE + + + type_synonym exc_code =" 4 bits " +datatype InterruptType = + I_U_Software + | I_S_Software + | I_M_Software + | I_U_Timer + | I_S_Timer + | I_M_Timer + | I_U_External + | I_S_External + | I_M_External + + + datatype ExceptionType = E_Fetch_Addr_Align | E_Fetch_Access_Fault @@ -90,16 +107,7 @@ datatype ExceptionType = -datatype InterruptType = - I_U_Software - | I_S_Software - | I_M_Software - | I_U_Timer - | I_S_Timer - | I_M_Timer - | I_U_External - | I_S_External - | I_M_External +datatype exception = Error_not_implemented " (string)" | Error_internal_error " (unit)" @@ -109,10 +117,6 @@ datatype TrapVectorMode = TV_Direct | TV_Vector | TV_Reserved -datatype exception = Error_not_implemented " (string)" | Error_internal_error " (unit)" - - - type_synonym ext_status =" 2 bits " datatype ExtStatus = Off | Initial | Clean | Dirty @@ -161,15 +165,15 @@ datatype ropw = RISCV_ADDW | RISCV_SUBW | RISCV_SLLW | RISCV_SRLW | RISCV_SRAW -datatype amoop = AMOSWAP | AMOADD | AMOXOR | AMOAND | AMOOR | AMOMIN | AMOMAX | AMOMINU | AMOMAXU +datatype sopw = RISCV_SLLIW | RISCV_SRLIW | RISCV_SRAIW -datatype csrop = CSRRW | CSRRS | CSRRC +datatype amoop = AMOSWAP | AMOADD | AMOXOR | AMOAND | AMOOR | AMOMIN | AMOMAX | AMOMINU | AMOMAXU -datatype word_width = BYTE | HALF | WORD | DOUBLE +datatype csrop = CSRRW | CSRRS | CSRRC @@ -330,6 +334,7 @@ datatype (plugins only: size) ast = | ADDIW " (( 12 bits * regbits * regbits))" | SHIFTW " (( 5 bits * regbits * regbits * sop))" | RTYPEW " ((regbits * regbits * regbits * ropw))" + | SHIFTIWOP " (( 5 bits * regbits * regbits * sopw))" | MUL " ((regbits * regbits * regbits * bool * bool * bool))" | DIV " ((regbits * regbits * regbits * bool))" | REM " ((regbits * regbits * regbits * bool))" |
